author | haftmann |
Thu, 19 Jun 2025 17:15:40 +0200 | |
changeset 82734 | 89347c0cc6a3 |
parent 82587 | 7415414bd9d8 |
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 |
81589 | 10 |
datatype 'a trrule = |
11 |
Parse_Rule of 'a * 'a | |
|
12 |
Print_Rule of 'a * 'a | |
|
13 |
Parse_Print_Rule of 'a * 'a |
|
42382
dcd983ee2c29
provide structure Syntax early (before structure Type), back-patch check/uncheck later;
wenzelm
parents:
42360
diff
changeset
|
14 |
type operations |
62897
8093203f0b89
prefer regular context update, to allow continuous editing of Pure;
wenzelm
parents:
62752
diff
changeset
|
15 |
val install_operations: operations -> theory -> theory |
42288
2074b31650e6
discontinued special treatment of structure Syntax_Ext (formerly Syn_Ext);
wenzelm
parents:
42287
diff
changeset
|
16 |
val root: string Config.T |
46512
4f9f61f9b535
simplified configuration options for syntax ambiguity;
wenzelm
parents:
46506
diff
changeset
|
17 |
val ambiguity_warning: bool Config.T |
42268 | 18 |
val ambiguity_limit: int Config.T |
59795 | 19 |
val encode_input: Input.source -> XML.tree |
20 |
val implode_input: Input.source -> string |
|
21 |
val read_input_pos: string -> Position.T |
|
22 |
val read_input: string -> Input.source |
|
23 |
val parse_input: Proof.context -> (XML.tree list -> 'a) -> |
|
55828
42ac3cfb89f6
clarified language markup: added "delimited" property;
wenzelm
parents:
52160
diff
changeset
|
24 |
(bool -> Markup.T) -> (Symbol_Pos.T list * Position.T -> 'a) -> string -> 'a |
81589 | 25 |
val parse_trrule: Proof.context -> (string * string) trrule -> Ast.ast trrule |
24263 | 26 |
val parse_sort: Proof.context -> string -> sort |
27 |
val parse_typ: Proof.context -> string -> typ |
|
28 |
val parse_term: Proof.context -> string -> term |
|
29 |
val parse_prop: Proof.context -> string -> term |
|
24768 | 30 |
val unparse_sort: Proof.context -> sort -> Pretty.T |
24923 | 31 |
val unparse_classrel: Proof.context -> class list -> Pretty.T |
32 |
val unparse_arity: Proof.context -> arity -> Pretty.T |
|
24768 | 33 |
val unparse_typ: Proof.context -> typ -> Pretty.T |
34 |
val unparse_term: Proof.context -> term -> Pretty.T |
|
24488 | 35 |
val check_sort: Proof.context -> sort -> sort |
24512 | 36 |
val check_typ: Proof.context -> typ -> typ |
37 |
val check_term: Proof.context -> term -> term |
|
38 |
val check_prop: Proof.context -> term -> term |
|
24488 | 39 |
val check_typs: Proof.context -> typ list -> typ list |
24263 | 40 |
val check_terms: Proof.context -> term list -> term list |
41 |
val check_props: Proof.context -> term list -> term list |
|
24923 | 42 |
val uncheck_sort: Proof.context -> sort -> sort |
43 |
val uncheck_arity: Proof.context -> arity -> arity |
|
44 |
val uncheck_classrel: Proof.context -> class list -> class list |
|
24768 | 45 |
val uncheck_typs: Proof.context -> typ list -> typ list |
46 |
val uncheck_terms: Proof.context -> term list -> term list |
|
24263 | 47 |
val read_sort: Proof.context -> string -> sort |
48 |
val read_typ: Proof.context -> string -> typ |
|
24488 | 49 |
val read_term: Proof.context -> string -> term |
50 |
val read_prop: Proof.context -> string -> term |
|
46989 | 51 |
val read_typs: Proof.context -> string list -> typ list |
24263 | 52 |
val read_terms: Proof.context -> string list -> term list |
53 |
val read_props: Proof.context -> string list -> term list |
|
24709 | 54 |
val read_sort_global: theory -> string -> sort |
55 |
val read_typ_global: theory -> string -> typ |
|
56 |
val read_term_global: theory -> string -> term |
|
57 |
val read_prop_global: theory -> string -> term |
|
24923 | 58 |
val pretty_term: Proof.context -> term -> Pretty.T |
59 |
val pretty_typ: Proof.context -> typ -> Pretty.T |
|
60 |
val pretty_sort: Proof.context -> sort -> Pretty.T |
|
61 |
val pretty_classrel: Proof.context -> class list -> Pretty.T |
|
62 |
val pretty_arity: Proof.context -> arity -> Pretty.T |
|
63 |
val string_of_term: Proof.context -> term -> string |
|
64 |
val string_of_typ: Proof.context -> typ -> string |
|
65 |
val string_of_sort: Proof.context -> sort -> string |
|
66 |
val string_of_classrel: Proof.context -> class list -> string |
|
67 |
val string_of_arity: Proof.context -> arity -> string |
|
26951
030e4a818b39
moved global pretty/string_of functions from Sign to Syntax;
wenzelm
parents:
26704
diff
changeset
|
68 |
val is_pretty_global: Proof.context -> bool |
030e4a818b39
moved global pretty/string_of functions from Sign to Syntax;
wenzelm
parents:
26704
diff
changeset
|
69 |
val set_pretty_global: bool -> Proof.context -> Proof.context |
030e4a818b39
moved global pretty/string_of functions from Sign to Syntax;
wenzelm
parents:
26704
diff
changeset
|
70 |
val init_pretty_global: theory -> Proof.context |
61262
7bd1eb4b056e
tuned signature: eliminated pointless type Context.pretty;
wenzelm
parents:
59841
diff
changeset
|
71 |
val init_pretty: Context.generic -> Proof.context |
26951
030e4a818b39
moved global pretty/string_of functions from Sign to Syntax;
wenzelm
parents:
26704
diff
changeset
|
72 |
val pretty_term_global: theory -> term -> Pretty.T |
030e4a818b39
moved global pretty/string_of functions from Sign to Syntax;
wenzelm
parents:
26704
diff
changeset
|
73 |
val pretty_typ_global: theory -> typ -> Pretty.T |
030e4a818b39
moved global pretty/string_of functions from Sign to Syntax;
wenzelm
parents:
26704
diff
changeset
|
74 |
val pretty_sort_global: theory -> sort -> Pretty.T |
030e4a818b39
moved global pretty/string_of functions from Sign to Syntax;
wenzelm
parents:
26704
diff
changeset
|
75 |
val string_of_term_global: theory -> term -> string |
030e4a818b39
moved global pretty/string_of functions from Sign to Syntax;
wenzelm
parents:
26704
diff
changeset
|
76 |
val string_of_typ_global: theory -> typ -> string |
030e4a818b39
moved global pretty/string_of functions from Sign to Syntax;
wenzelm
parents:
26704
diff
changeset
|
77 |
val string_of_sort_global: theory -> sort -> string |
76082 | 78 |
val pretty_flexpair: Proof.context -> term * term -> Pretty.T list |
39135
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
79 |
type syntax |
80073
40f5ddeda2b4
further performance tuning (after f906f7f83dae): interactive mode is closer to earlier approach with Lazy.value, which could be relevant with rather complex grammars under tight memory situations;
wenzelm
parents:
78013
diff
changeset
|
80 |
val cache_persistent: bool Unsynchronized.ref |
40f5ddeda2b4
further performance tuning (after f906f7f83dae): interactive mode is closer to earlier approach with Lazy.value, which could be relevant with rather complex grammars under tight memory situations;
wenzelm
parents:
78013
diff
changeset
|
81 |
val cache_syntax: syntax -> unit |
39135
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
82 |
val eq_syntax: syntax * syntax -> bool |
69077
11529ae45786
more approximative prefix syntax, including binder;
wenzelm
parents:
69071
diff
changeset
|
83 |
datatype approx = Prefix of string | Infix of {assoc: Printer.assoc, delim: string, pri: int} |
11529ae45786
more approximative prefix syntax, including binder;
wenzelm
parents:
69071
diff
changeset
|
84 |
val get_approx: syntax -> string -> approx option |
80748
43c4817375bf
support for syntax const dependencies, with minimal integrity checks;
wenzelm
parents:
80742
diff
changeset
|
85 |
val get_consts: syntax -> string -> string list |
80739 | 86 |
val is_const: syntax -> string -> bool |
39135
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
87 |
val is_keyword: syntax -> string -> bool |
81601 | 88 |
val tokenize: syntax -> {raw: bool} -> Symbol_Pos.T list -> Lexicon.token list |
80991 | 89 |
val parse: syntax -> string -> Lexicon.token list -> Parser.tree list |
42253 | 90 |
val parse_ast_translation: syntax -> string -> (Proof.context -> Ast.ast list -> Ast.ast) option |
42255 | 91 |
val parse_rules: syntax -> string -> (Ast.ast * Ast.ast) list |
42253 | 92 |
val parse_translation: syntax -> string -> (Proof.context -> term list -> term) option |
42254 | 93 |
val print_translation: syntax -> string -> |
94 |
Proof.context -> typ -> term list -> term (*exception Match*) |
|
42255 | 95 |
val print_rules: syntax -> string -> (Ast.ast * Ast.ast) list |
42254 | 96 |
val print_ast_translation: syntax -> string -> |
97 |
Proof.context -> Ast.ast list -> Ast.ast (*exception Match*) |
|
42255 | 98 |
val prtabs: syntax -> Printer.prtabs |
81176
c0522b2d3df6
clarified signature: more explicit Syntax.print_mode_tabs, depending on print_mode_value ();
wenzelm
parents:
81163
diff
changeset
|
99 |
val print_mode_tabs: syntax -> Printer.prtab list |
39135
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
100 |
type mode |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
101 |
val mode_default: mode |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
102 |
val mode_input: mode |
42294 | 103 |
val empty_syntax: syntax |
45632
b23c42b9f78a
prefer Parser.make_gram over Parser.merge_gram, to approximate n-ary merges on theory import;
wenzelm
parents:
45490
diff
changeset
|
104 |
val merge_syntax: syntax * syntax -> syntax |
39135
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
105 |
val print_gram: syntax -> unit |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
106 |
val print_trans: syntax -> unit |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
107 |
val print_syntax: syntax -> unit |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
108 |
val map_trrule: ('a -> 'b) -> 'a trrule -> 'b trrule |
80897
5328d67ec647
more explicit context for syn_ext/mixfix operations, but it often degenerates to background theory;
wenzelm
parents:
80767
diff
changeset
|
109 |
val update_trfuns: Proof.context -> |
42224
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
wenzelm
parents:
42223
diff
changeset
|
110 |
(string * ((Proof.context -> Ast.ast list -> Ast.ast) * stamp)) list * |
39135
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
111 |
(string * ((Proof.context -> term list -> term) * stamp)) list * |
42247
12fe41a92cd5
typed_print_translation: discontinued show_sorts argument;
wenzelm
parents:
42245
diff
changeset
|
112 |
(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:
42223
diff
changeset
|
113 |
(string * ((Proof.context -> Ast.ast list -> Ast.ast) * stamp)) list -> syntax -> syntax |
80897
5328d67ec647
more explicit context for syn_ext/mixfix operations, but it often degenerates to background theory;
wenzelm
parents:
80767
diff
changeset
|
114 |
val update_type_gram: Proof.context -> bool -> mode -> (string * typ * mixfix) list -> |
5328d67ec647
more explicit context for syn_ext/mixfix operations, but it often degenerates to background theory;
wenzelm
parents:
80767
diff
changeset
|
115 |
syntax -> syntax |
5328d67ec647
more explicit context for syn_ext/mixfix operations, but it often degenerates to background theory;
wenzelm
parents:
80767
diff
changeset
|
116 |
val update_const_gram: Proof.context -> bool -> string list -> |
39135
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
117 |
mode -> (string * typ * mixfix) list -> syntax -> syntax |
80897
5328d67ec647
more explicit context for syn_ext/mixfix operations, but it often degenerates to background theory;
wenzelm
parents:
80767
diff
changeset
|
118 |
val update_consts: Proof.context -> (string * string list) list -> syntax -> syntax |
81590 | 119 |
val update_trrules: Proof.context -> bool -> Ast.ast trrule list -> syntax -> syntax |
81116
0fb1e2dd4122
first-class support for "unbundle no foobar_syntax" -- avoid redundant "bundle no_foobar_syntax" definitions;
wenzelm
parents:
81005
diff
changeset
|
120 |
val get_polarity: Proof.context -> bool |
0fb1e2dd4122
first-class support for "unbundle no foobar_syntax" -- avoid redundant "bundle no_foobar_syntax" definitions;
wenzelm
parents:
81005
diff
changeset
|
121 |
val set_polarity: bool -> Proof.context -> Proof.context |
0fb1e2dd4122
first-class support for "unbundle no foobar_syntax" -- avoid redundant "bundle no_foobar_syntax" definitions;
wenzelm
parents:
81005
diff
changeset
|
122 |
val set_polarity_generic: bool -> Context.generic -> Context.generic |
0fb1e2dd4122
first-class support for "unbundle no foobar_syntax" -- avoid redundant "bundle no_foobar_syntax" definitions;
wenzelm
parents:
81005
diff
changeset
|
123 |
val effective_polarity: Proof.context -> bool -> bool |
81993 | 124 |
val effective_polarity_global: theory -> bool -> bool |
125 |
val effective_polarity_generic: Context.generic -> bool -> bool |
|
42476 | 126 |
val const: string -> term |
127 |
val free: string -> term |
|
128 |
val var: indexname -> term |
|
2383 | 129 |
end; |
0 | 130 |
|
12094 | 131 |
structure Syntax: SYNTAX = |
0 | 132 |
struct |
133 |
||
42288
2074b31650e6
discontinued special treatment of structure Syntax_Ext (formerly Syn_Ext);
wenzelm
parents:
42287
diff
changeset
|
134 |
|
39135
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
135 |
(** inner syntax operations **) |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
136 |
|
62897
8093203f0b89
prefer regular context update, to allow continuous editing of Pure;
wenzelm
parents:
62752
diff
changeset
|
137 |
(* operations *) |
42382
dcd983ee2c29
provide structure Syntax early (before structure Type), back-patch check/uncheck later;
wenzelm
parents:
42360
diff
changeset
|
138 |
|
81589 | 139 |
datatype 'a trrule = |
140 |
Parse_Rule of 'a * 'a | |
|
141 |
Print_Rule of 'a * 'a | |
|
142 |
Parse_Print_Rule of 'a * 'a; |
|
143 |
||
42382
dcd983ee2c29
provide structure Syntax early (before structure Type), back-patch check/uncheck later;
wenzelm
parents:
42360
diff
changeset
|
144 |
type operations = |
81589 | 145 |
{parse_trrule: Proof.context -> (string * string) trrule -> Ast.ast trrule, |
146 |
parse_sort: Proof.context -> string -> sort, |
|
42382
dcd983ee2c29
provide structure Syntax early (before structure Type), back-patch check/uncheck later;
wenzelm
parents:
42360
diff
changeset
|
147 |
parse_typ: Proof.context -> string -> typ, |
dcd983ee2c29
provide structure Syntax early (before structure Type), back-patch check/uncheck later;
wenzelm
parents:
42360
diff
changeset
|
148 |
parse_term: Proof.context -> string -> term, |
dcd983ee2c29
provide structure Syntax early (before structure Type), back-patch check/uncheck later;
wenzelm
parents:
42360
diff
changeset
|
149 |
parse_prop: Proof.context -> string -> term, |
dcd983ee2c29
provide structure Syntax early (before structure Type), back-patch check/uncheck later;
wenzelm
parents:
42360
diff
changeset
|
150 |
unparse_sort: Proof.context -> sort -> Pretty.T, |
dcd983ee2c29
provide structure Syntax early (before structure Type), back-patch check/uncheck later;
wenzelm
parents:
42360
diff
changeset
|
151 |
unparse_typ: Proof.context -> typ -> Pretty.T, |
dcd983ee2c29
provide structure Syntax early (before structure Type), back-patch check/uncheck later;
wenzelm
parents:
42360
diff
changeset
|
152 |
unparse_term: Proof.context -> term -> Pretty.T, |
dcd983ee2c29
provide structure Syntax early (before structure Type), back-patch check/uncheck later;
wenzelm
parents:
42360
diff
changeset
|
153 |
check_typs: Proof.context -> typ list -> typ list, |
dcd983ee2c29
provide structure Syntax early (before structure Type), back-patch check/uncheck later;
wenzelm
parents:
42360
diff
changeset
|
154 |
check_terms: Proof.context -> term list -> term list, |
dcd983ee2c29
provide structure Syntax early (before structure Type), back-patch check/uncheck later;
wenzelm
parents:
42360
diff
changeset
|
155 |
check_props: Proof.context -> term list -> term list, |
dcd983ee2c29
provide structure Syntax early (before structure Type), back-patch check/uncheck later;
wenzelm
parents:
42360
diff
changeset
|
156 |
uncheck_typs: Proof.context -> typ list -> typ list, |
dcd983ee2c29
provide structure Syntax early (before structure Type), back-patch check/uncheck later;
wenzelm
parents:
42360
diff
changeset
|
157 |
uncheck_terms: Proof.context -> term list -> term list}; |
dcd983ee2c29
provide structure Syntax early (before structure Type), back-patch check/uncheck later;
wenzelm
parents:
42360
diff
changeset
|
158 |
|
62897
8093203f0b89
prefer regular context update, to allow continuous editing of Pure;
wenzelm
parents:
62752
diff
changeset
|
159 |
structure Operations = Theory_Data |
8093203f0b89
prefer regular context update, to allow continuous editing of Pure;
wenzelm
parents:
62752
diff
changeset
|
160 |
( |
8093203f0b89
prefer regular context update, to allow continuous editing of Pure;
wenzelm
parents:
62752
diff
changeset
|
161 |
type T = operations option; |
8093203f0b89
prefer regular context update, to allow continuous editing of Pure;
wenzelm
parents:
62752
diff
changeset
|
162 |
val empty = NONE; |
8093203f0b89
prefer regular context update, to allow continuous editing of Pure;
wenzelm
parents:
62752
diff
changeset
|
163 |
val merge = merge_options; |
8093203f0b89
prefer regular context update, to allow continuous editing of Pure;
wenzelm
parents:
62752
diff
changeset
|
164 |
); |
8093203f0b89
prefer regular context update, to allow continuous editing of Pure;
wenzelm
parents:
62752
diff
changeset
|
165 |
|
8093203f0b89
prefer regular context update, to allow continuous editing of Pure;
wenzelm
parents:
62752
diff
changeset
|
166 |
fun install_operations ops = |
8093203f0b89
prefer regular context update, to allow continuous editing of Pure;
wenzelm
parents:
62752
diff
changeset
|
167 |
Operations.map |
8093203f0b89
prefer regular context update, to allow continuous editing of Pure;
wenzelm
parents:
62752
diff
changeset
|
168 |
(fn NONE => SOME ops |
8093203f0b89
prefer regular context update, to allow continuous editing of Pure;
wenzelm
parents:
62752
diff
changeset
|
169 |
| SOME _ => raise Fail "Inner syntax operations already installed"); |
42382
dcd983ee2c29
provide structure Syntax early (before structure Type), back-patch check/uncheck later;
wenzelm
parents:
42360
diff
changeset
|
170 |
|
dcd983ee2c29
provide structure Syntax early (before structure Type), back-patch check/uncheck later;
wenzelm
parents:
42360
diff
changeset
|
171 |
fun operation which ctxt x = |
62897
8093203f0b89
prefer regular context update, to allow continuous editing of Pure;
wenzelm
parents:
62752
diff
changeset
|
172 |
(case Operations.get (Proof_Context.theory_of ctxt) of |
42382
dcd983ee2c29
provide structure Syntax early (before structure Type), back-patch check/uncheck later;
wenzelm
parents:
42360
diff
changeset
|
173 |
NONE => raise Fail "Inner syntax operations not installed" |
dcd983ee2c29
provide structure Syntax early (before structure Type), back-patch check/uncheck later;
wenzelm
parents:
42360
diff
changeset
|
174 |
| SOME ops => which ops ctxt x); |
dcd983ee2c29
provide structure Syntax early (before structure Type), back-patch check/uncheck later;
wenzelm
parents:
42360
diff
changeset
|
175 |
|
dcd983ee2c29
provide structure Syntax early (before structure Type), back-patch check/uncheck later;
wenzelm
parents:
42360
diff
changeset
|
176 |
|
42268 | 177 |
(* configuration options *) |
178 |
||
69575 | 179 |
val root = Config.declare_string ("syntax_root", \<^here>) (K "any"); |
180 |
val ambiguity_warning = Config.declare_bool ("syntax_ambiguity_warning", \<^here>) (K true); |
|
181 |
val ambiguity_limit = Config.declare_int ("syntax_ambiguity_limit", \<^here>) (K 10); |
|
44069 | 182 |
|
42268 | 183 |
|
59795 | 184 |
(* formal input *) |
185 |
||
186 |
fun encode_input source = |
|
187 |
let |
|
188 |
val delimited = Input.is_delimited source; |
|
189 |
val pos = Input.pos_of source; |
|
190 |
val text = Input.text_of source; |
|
191 |
in XML.Elem (Markup.input delimited (Position.properties_of pos), [XML.Text text]) end; |
|
192 |
||
193 |
val implode_input = encode_input #> YXML.string_of; |
|
39135
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
194 |
|
55828
42ac3cfb89f6
clarified language markup: added "delimited" property;
wenzelm
parents:
52160
diff
changeset
|
195 |
local |
42ac3cfb89f6
clarified language markup: added "delimited" property;
wenzelm
parents:
52160
diff
changeset
|
196 |
|
59795 | 197 |
fun input_range (XML.Elem ((name, props), _)) = |
198 |
if name = Markup.inputN |
|
58978
e42da880c61e
more position information, e.g. relevant for errors in generated ML source;
wenzelm
parents:
58047
diff
changeset
|
199 |
then (Markup.is_delimited props, Position.range_of_properties props) |
e42da880c61e
more position information, e.g. relevant for errors in generated ML source;
wenzelm
parents:
58047
diff
changeset
|
200 |
else (false, Position.no_range) |
59795 | 201 |
| input_range (XML.Text _) = (false, Position.no_range); |
46849 | 202 |
|
59795 | 203 |
fun input_source tree = |
39135
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
204 |
let |
39555
ccb223a4d49c
added XML.content_of convenience -- cover XML.body, which is the general situation;
wenzelm
parents:
39510
diff
changeset
|
205 |
val text = XML.content_of [tree]; |
59795 | 206 |
val (delimited, range) = input_range tree; |
59064 | 207 |
in Input.source delimited text range end; |
55828
42ac3cfb89f6
clarified language markup: added "delimited" property;
wenzelm
parents:
52160
diff
changeset
|
208 |
|
42ac3cfb89f6
clarified language markup: added "delimited" property;
wenzelm
parents:
52160
diff
changeset
|
209 |
in |
42ac3cfb89f6
clarified language markup: added "delimited" property;
wenzelm
parents:
52160
diff
changeset
|
210 |
|
59795 | 211 |
fun read_input_pos str = #1 (#2 (input_range (YXML.parse str handle Fail msg => error msg))); |
39135
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
212 |
|
59795 | 213 |
fun read_input str = input_source (YXML.parse str handle Fail msg => error msg); |
43731
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
wenzelm
parents:
43558
diff
changeset
|
214 |
|
59795 | 215 |
fun parse_input ctxt decode markup parse str = |
43731
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
wenzelm
parents:
43558
diff
changeset
|
216 |
let |
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
wenzelm
parents:
43558
diff
changeset
|
217 |
fun parse_tree tree = |
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
wenzelm
parents:
43558
diff
changeset
|
218 |
let |
59795 | 219 |
val source = input_source tree; |
59064 | 220 |
val syms = Input.source_explode source; |
221 |
val pos = Input.pos_of source; |
|
69971 | 222 |
val _ = |
223 |
Context_Position.reports ctxt |
|
224 |
[(pos, Markup.cartouche), (pos, markup (Input.is_delimited source))]; |
|
69589
e15f053a42d8
support for isabelle update -u inner_syntax_cartouches;
wenzelm
parents:
69576
diff
changeset
|
225 |
val _ = |
e15f053a42d8
support for isabelle update -u inner_syntax_cartouches;
wenzelm
parents:
69576
diff
changeset
|
226 |
if Options.default_bool "update_inner_syntax_cartouches" then |
e15f053a42d8
support for isabelle update -u inner_syntax_cartouches;
wenzelm
parents:
69576
diff
changeset
|
227 |
Context_Position.report_text ctxt |
e15f053a42d8
support for isabelle update -u inner_syntax_cartouches;
wenzelm
parents:
69576
diff
changeset
|
228 |
pos Markup.update (cartouche (Symbol_Pos.content syms)) |
e15f053a42d8
support for isabelle update -u inner_syntax_cartouches;
wenzelm
parents:
69576
diff
changeset
|
229 |
else (); |
58978
e42da880c61e
more position information, e.g. relevant for errors in generated ML source;
wenzelm
parents:
58047
diff
changeset
|
230 |
in parse (syms, pos) end; |
43731
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
wenzelm
parents:
43558
diff
changeset
|
231 |
in |
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
wenzelm
parents:
43558
diff
changeset
|
232 |
(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:
43558
diff
changeset
|
233 |
body as [tree as XML.Elem ((name, _), _)] => |
59795 | 234 |
if name = Markup.inputN then parse_tree tree else decode body |
43731
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
wenzelm
parents:
43558
diff
changeset
|
235 |
| [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:
43558
diff
changeset
|
236 |
| body => decode body) |
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
wenzelm
parents:
43558
diff
changeset
|
237 |
end; |
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
wenzelm
parents:
43558
diff
changeset
|
238 |
|
55828
42ac3cfb89f6
clarified language markup: added "delimited" property;
wenzelm
parents:
52160
diff
changeset
|
239 |
end; |
42ac3cfb89f6
clarified language markup: added "delimited" property;
wenzelm
parents:
52160
diff
changeset
|
240 |
|
39135
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
241 |
|
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
242 |
(* (un)parsing *) |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
243 |
|
81589 | 244 |
val parse_trrule = operation #parse_trrule; |
39135
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
245 |
val parse_sort = operation #parse_sort; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
246 |
val parse_typ = operation #parse_typ; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
247 |
val parse_term = operation #parse_term; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
248 |
val parse_prop = operation #parse_prop; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
249 |
val unparse_sort = operation #unparse_sort; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
250 |
val unparse_typ = operation #unparse_typ; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
251 |
val unparse_term = operation #unparse_term; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
252 |
|
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
253 |
|
45429 | 254 |
(* (un)checking *) |
42382
dcd983ee2c29
provide structure Syntax early (before structure Type), back-patch check/uncheck later;
wenzelm
parents:
42360
diff
changeset
|
255 |
|
dcd983ee2c29
provide structure Syntax early (before structure Type), back-patch check/uncheck later;
wenzelm
parents:
42360
diff
changeset
|
256 |
val check_typs = operation #check_typs; |
dcd983ee2c29
provide structure Syntax early (before structure Type), back-patch check/uncheck later;
wenzelm
parents:
42360
diff
changeset
|
257 |
val check_terms = operation #check_terms; |
dcd983ee2c29
provide structure Syntax early (before structure Type), back-patch check/uncheck later;
wenzelm
parents:
42360
diff
changeset
|
258 |
val check_props = operation #check_props; |
39135
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
259 |
|
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
260 |
val check_typ = singleton o check_typs; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
261 |
val check_term = singleton o check_terms; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
262 |
val check_prop = singleton o check_props; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
263 |
|
42382
dcd983ee2c29
provide structure Syntax early (before structure Type), back-patch check/uncheck later;
wenzelm
parents:
42360
diff
changeset
|
264 |
val uncheck_typs = operation #uncheck_typs; |
dcd983ee2c29
provide structure Syntax early (before structure Type), back-patch check/uncheck later;
wenzelm
parents:
42360
diff
changeset
|
265 |
val uncheck_terms = operation #uncheck_terms; |
42402
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
wenzelm
parents:
42383
diff
changeset
|
266 |
|
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
wenzelm
parents:
42383
diff
changeset
|
267 |
|
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
wenzelm
parents:
42383
diff
changeset
|
268 |
(* derived operations for algebra of sorts *) |
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
wenzelm
parents:
42383
diff
changeset
|
269 |
|
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
wenzelm
parents:
42383
diff
changeset
|
270 |
local |
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
wenzelm
parents:
42383
diff
changeset
|
271 |
|
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
wenzelm
parents:
42383
diff
changeset
|
272 |
fun map_sort f S = |
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
wenzelm
parents:
42383
diff
changeset
|
273 |
(case f (TFree ("", S)) of |
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
wenzelm
parents:
42383
diff
changeset
|
274 |
TFree ("", S') => S' |
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
wenzelm
parents:
42383
diff
changeset
|
275 |
| _ => raise TYPE ("map_sort", [TFree ("", S)], [])); |
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
wenzelm
parents:
42383
diff
changeset
|
276 |
|
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
wenzelm
parents:
42383
diff
changeset
|
277 |
in |
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
wenzelm
parents:
42383
diff
changeset
|
278 |
|
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
wenzelm
parents:
42383
diff
changeset
|
279 |
val check_sort = map_sort o check_typ; |
39135
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
280 |
val uncheck_sort = map_sort o singleton o uncheck_typs; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
281 |
|
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
282 |
end; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
283 |
|
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
284 |
|
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
285 |
val uncheck_classrel = map o singleton o uncheck_sort; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
286 |
|
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
287 |
fun unparse_classrel ctxt cs = Pretty.block (flat |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
288 |
(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:
39126
diff
changeset
|
289 |
|
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
290 |
fun uncheck_arity ctxt (a, Ss, S) = |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
291 |
let |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
292 |
val T = Type (a, replicate (length Ss) dummyT); |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
293 |
val a' = |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
294 |
(case singleton (uncheck_typs ctxt) T of |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
295 |
Type (a', _) => a' |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
296 |
| T => raise TYPE ("uncheck_arity", [T], [])); |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
297 |
val Ss' = map (uncheck_sort ctxt) Ss; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
298 |
val S' = uncheck_sort ctxt S; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
299 |
in (a', Ss', S') end; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
300 |
|
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
301 |
fun unparse_arity ctxt (a, Ss, S) = |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
302 |
let |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
303 |
val prtT = unparse_typ ctxt (Type (a, [])); |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
304 |
val dom = |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
305 |
if null Ss then [] |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
306 |
else [Pretty.list "(" ")" (map (unparse_sort ctxt) Ss), Pretty.brk 1]; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
307 |
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:
39126
diff
changeset
|
308 |
|
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
309 |
|
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
310 |
(* read = parse + check *) |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
311 |
|
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
312 |
fun read_sort ctxt = parse_sort ctxt #> check_sort ctxt; |
46989 | 313 |
|
314 |
fun read_typs ctxt = |
|
58993
302104d8366b
prefer independent parallel map where user input is processed -- avoid non-deterministic feedback in error situations;
wenzelm
parents:
58978
diff
changeset
|
315 |
grouped 10 Par_List.map_independent (parse_typ ctxt) #> check_typs ctxt; |
39135
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
316 |
|
46989 | 317 |
fun read_terms ctxt = |
58993
302104d8366b
prefer independent parallel map where user input is processed -- avoid non-deterministic feedback in error situations;
wenzelm
parents:
58978
diff
changeset
|
318 |
grouped 10 Par_List.map_independent (parse_term ctxt) #> check_terms ctxt; |
39135
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
319 |
|
46989 | 320 |
fun read_props ctxt = |
58993
302104d8366b
prefer independent parallel map where user input is processed -- avoid non-deterministic feedback in error situations;
wenzelm
parents:
58978
diff
changeset
|
321 |
grouped 10 Par_List.map_independent (parse_prop ctxt) #> check_props ctxt; |
46989 | 322 |
|
323 |
val read_typ = singleton o read_typs; |
|
39135
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
324 |
val read_term = singleton o read_terms; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
325 |
val read_prop = singleton o read_props; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
326 |
|
42360 | 327 |
val read_sort_global = read_sort o Proof_Context.init_global; |
328 |
val read_typ_global = read_typ o Proof_Context.init_global; |
|
329 |
val read_term_global = read_term o Proof_Context.init_global; |
|
330 |
val read_prop_global = read_prop o Proof_Context.init_global; |
|
39135
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
331 |
|
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
332 |
|
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
333 |
(* pretty = uncheck + unparse *) |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
334 |
|
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
335 |
fun pretty_term ctxt = singleton (uncheck_terms ctxt) #> unparse_term ctxt; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
336 |
fun pretty_typ ctxt = singleton (uncheck_typs ctxt) #> unparse_typ ctxt; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
337 |
fun pretty_sort ctxt = uncheck_sort ctxt #> unparse_sort ctxt; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
338 |
fun pretty_classrel ctxt = uncheck_classrel ctxt #> unparse_classrel ctxt; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
339 |
fun pretty_arity ctxt = uncheck_arity ctxt #> unparse_arity ctxt; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
340 |
|
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
341 |
val string_of_term = Pretty.string_of oo pretty_term; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
342 |
val string_of_typ = Pretty.string_of oo pretty_typ; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
343 |
val string_of_sort = Pretty.string_of oo pretty_sort; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
344 |
val string_of_classrel = Pretty.string_of oo pretty_classrel; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
345 |
val string_of_arity = Pretty.string_of oo pretty_arity; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
346 |
|
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
347 |
|
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
348 |
(* global pretty printing *) |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
349 |
|
69575 | 350 |
val pretty_global = Config.declare_bool ("Syntax.pretty_global", \<^here>) (K false); |
69576 | 351 |
val is_pretty_global = Config.apply pretty_global; |
39508
dfacdb01e1ec
simplified some internal flags using Config.T instead of full-blown Proof_Data;
wenzelm
parents:
39507
diff
changeset
|
352 |
val set_pretty_global = Config.put pretty_global; |
42360 | 353 |
val init_pretty_global = set_pretty_global true o Proof_Context.init_global; |
61262
7bd1eb4b056e
tuned signature: eliminated pointless type Context.pretty;
wenzelm
parents:
59841
diff
changeset
|
354 |
val init_pretty = Context.cases init_pretty_global I; |
39135
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
355 |
|
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
356 |
val pretty_term_global = pretty_term o init_pretty_global; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
357 |
val pretty_typ_global = pretty_typ o init_pretty_global; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
358 |
val pretty_sort_global = pretty_sort o init_pretty_global; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
359 |
|
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
360 |
val string_of_term_global = string_of_term o init_pretty_global; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
361 |
val string_of_typ_global = string_of_typ o init_pretty_global; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
362 |
val string_of_sort_global = string_of_sort o init_pretty_global; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
363 |
|
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
364 |
|
70443 | 365 |
(* derived operations *) |
366 |
||
76082 | 367 |
fun pretty_flexpair ctxt (t, u) = |
70443 | 368 |
[pretty_term ctxt t, Pretty.str " \<equiv>\<^sup>?", Pretty.brk 1, pretty_term ctxt u]; |
369 |
||
370 |
||
39135
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
371 |
|
81588
81a72b7fcb0c
minor performance tuning: avoid excessive Library.insert and Symbol.explode operations;
wenzelm
parents:
81176
diff
changeset
|
372 |
(** keywords **) |
81a72b7fcb0c
minor performance tuning: avoid excessive Library.insert and Symbol.explode operations;
wenzelm
parents:
81176
diff
changeset
|
373 |
|
81a72b7fcb0c
minor performance tuning: avoid excessive Library.insert and Symbol.explode operations;
wenzelm
parents:
81176
diff
changeset
|
374 |
abstype keywords = Keywords of Symset.T * Scan.lexicon lazy |
81a72b7fcb0c
minor performance tuning: avoid excessive Library.insert and Symbol.explode operations;
wenzelm
parents:
81176
diff
changeset
|
375 |
with |
81a72b7fcb0c
minor performance tuning: avoid excessive Library.insert and Symbol.explode operations;
wenzelm
parents:
81176
diff
changeset
|
376 |
|
81a72b7fcb0c
minor performance tuning: avoid excessive Library.insert and Symbol.explode operations;
wenzelm
parents:
81176
diff
changeset
|
377 |
val empty_keywords = |
81a72b7fcb0c
minor performance tuning: avoid excessive Library.insert and Symbol.explode operations;
wenzelm
parents:
81176
diff
changeset
|
378 |
Keywords (Symset.empty, Lazy.value Scan.empty_lexicon); |
81a72b7fcb0c
minor performance tuning: avoid excessive Library.insert and Symbol.explode operations;
wenzelm
parents:
81176
diff
changeset
|
379 |
|
81a72b7fcb0c
minor performance tuning: avoid excessive Library.insert and Symbol.explode operations;
wenzelm
parents:
81176
diff
changeset
|
380 |
fun make_keywords set = |
81a72b7fcb0c
minor performance tuning: avoid excessive Library.insert and Symbol.explode operations;
wenzelm
parents:
81176
diff
changeset
|
381 |
let fun lex () = Scan.build_lexicon (set |> Symset.fold (Scan.extend_lexicon o Symbol.explode)) |
81a72b7fcb0c
minor performance tuning: avoid excessive Library.insert and Symbol.explode operations;
wenzelm
parents:
81176
diff
changeset
|
382 |
in Keywords (set, Lazy.lazy lex) end; |
81a72b7fcb0c
minor performance tuning: avoid excessive Library.insert and Symbol.explode operations;
wenzelm
parents:
81176
diff
changeset
|
383 |
|
81a72b7fcb0c
minor performance tuning: avoid excessive Library.insert and Symbol.explode operations;
wenzelm
parents:
81176
diff
changeset
|
384 |
fun add_keywords s (keywords as Keywords (set, lex)) = |
81a72b7fcb0c
minor performance tuning: avoid excessive Library.insert and Symbol.explode operations;
wenzelm
parents:
81176
diff
changeset
|
385 |
if Symset.member set s then keywords |
81a72b7fcb0c
minor performance tuning: avoid excessive Library.insert and Symbol.explode operations;
wenzelm
parents:
81176
diff
changeset
|
386 |
else |
81a72b7fcb0c
minor performance tuning: avoid excessive Library.insert and Symbol.explode operations;
wenzelm
parents:
81176
diff
changeset
|
387 |
let |
81a72b7fcb0c
minor performance tuning: avoid excessive Library.insert and Symbol.explode operations;
wenzelm
parents:
81176
diff
changeset
|
388 |
val set' = Symset.insert s set; |
81a72b7fcb0c
minor performance tuning: avoid excessive Library.insert and Symbol.explode operations;
wenzelm
parents:
81176
diff
changeset
|
389 |
val lex' = Lazy.map_finished (fn x => Scan.extend_lexicon (Symbol.explode s) x) lex; |
81a72b7fcb0c
minor performance tuning: avoid excessive Library.insert and Symbol.explode operations;
wenzelm
parents:
81176
diff
changeset
|
390 |
in Keywords (set', lex') end; |
81a72b7fcb0c
minor performance tuning: avoid excessive Library.insert and Symbol.explode operations;
wenzelm
parents:
81176
diff
changeset
|
391 |
|
81a72b7fcb0c
minor performance tuning: avoid excessive Library.insert and Symbol.explode operations;
wenzelm
parents:
81176
diff
changeset
|
392 |
fun merge_keywords (keywords1 as Keywords (set1, lex1), keywords2 as Keywords (set2, lex2)) = |
81a72b7fcb0c
minor performance tuning: avoid excessive Library.insert and Symbol.explode operations;
wenzelm
parents:
81176
diff
changeset
|
393 |
if pointer_eq (keywords1, keywords2) then keywords1 |
81a72b7fcb0c
minor performance tuning: avoid excessive Library.insert and Symbol.explode operations;
wenzelm
parents:
81176
diff
changeset
|
394 |
else if Symset.is_empty set1 then keywords2 |
81a72b7fcb0c
minor performance tuning: avoid excessive Library.insert and Symbol.explode operations;
wenzelm
parents:
81176
diff
changeset
|
395 |
else if Symset.is_empty set2 then keywords1 |
81a72b7fcb0c
minor performance tuning: avoid excessive Library.insert and Symbol.explode operations;
wenzelm
parents:
81176
diff
changeset
|
396 |
else if Symset.subset (set1, set2) then keywords2 |
81a72b7fcb0c
minor performance tuning: avoid excessive Library.insert and Symbol.explode operations;
wenzelm
parents:
81176
diff
changeset
|
397 |
else if Symset.subset (set2, set1) then keywords1 |
81a72b7fcb0c
minor performance tuning: avoid excessive Library.insert and Symbol.explode operations;
wenzelm
parents:
81176
diff
changeset
|
398 |
else |
81a72b7fcb0c
minor performance tuning: avoid excessive Library.insert and Symbol.explode operations;
wenzelm
parents:
81176
diff
changeset
|
399 |
let |
81a72b7fcb0c
minor performance tuning: avoid excessive Library.insert and Symbol.explode operations;
wenzelm
parents:
81176
diff
changeset
|
400 |
val set' = Symset.merge (set1, set2); |
81a72b7fcb0c
minor performance tuning: avoid excessive Library.insert and Symbol.explode operations;
wenzelm
parents:
81176
diff
changeset
|
401 |
val lex' = Lazy.value (Scan.merge_lexicons (apply2 Lazy.force (lex1, lex2))); |
81a72b7fcb0c
minor performance tuning: avoid excessive Library.insert and Symbol.explode operations;
wenzelm
parents:
81176
diff
changeset
|
402 |
in Keywords (set', lex') end; |
81a72b7fcb0c
minor performance tuning: avoid excessive Library.insert and Symbol.explode operations;
wenzelm
parents:
81176
diff
changeset
|
403 |
|
81a72b7fcb0c
minor performance tuning: avoid excessive Library.insert and Symbol.explode operations;
wenzelm
parents:
81176
diff
changeset
|
404 |
fun member_keywords (Keywords (set, _)) = Symset.member set; |
81a72b7fcb0c
minor performance tuning: avoid excessive Library.insert and Symbol.explode operations;
wenzelm
parents:
81176
diff
changeset
|
405 |
|
81a72b7fcb0c
minor performance tuning: avoid excessive Library.insert and Symbol.explode operations;
wenzelm
parents:
81176
diff
changeset
|
406 |
fun dest_keywords (Keywords (set, _)) = sort_strings (Symset.dest set); |
81a72b7fcb0c
minor performance tuning: avoid excessive Library.insert and Symbol.explode operations;
wenzelm
parents:
81176
diff
changeset
|
407 |
|
81a72b7fcb0c
minor performance tuning: avoid excessive Library.insert and Symbol.explode operations;
wenzelm
parents:
81176
diff
changeset
|
408 |
fun tokenize_keywords (Keywords (_, lex)) = Lexicon.tokenize (Lazy.force lex); |
81a72b7fcb0c
minor performance tuning: avoid excessive Library.insert and Symbol.explode operations;
wenzelm
parents:
81176
diff
changeset
|
409 |
|
81a72b7fcb0c
minor performance tuning: avoid excessive Library.insert and Symbol.explode operations;
wenzelm
parents:
81176
diff
changeset
|
410 |
end; |
81a72b7fcb0c
minor performance tuning: avoid excessive Library.insert and Symbol.explode operations;
wenzelm
parents:
81176
diff
changeset
|
411 |
|
81a72b7fcb0c
minor performance tuning: avoid excessive Library.insert and Symbol.explode operations;
wenzelm
parents:
81176
diff
changeset
|
412 |
|
81a72b7fcb0c
minor performance tuning: avoid excessive Library.insert and Symbol.explode operations;
wenzelm
parents:
81176
diff
changeset
|
413 |
|
237
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
175
diff
changeset
|
414 |
(** tables of translation functions **) |
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
175
diff
changeset
|
415 |
|
5692 | 416 |
(* parse (ast) translations *) |
417 |
||
23655
d2d1138e0ddc
replaced exception TableFun/GraphFun.DUPS by TableFun/GraphFun.DUP;
wenzelm
parents:
23631
diff
changeset
|
418 |
fun err_dup_trfun name c = |
d2d1138e0ddc
replaced exception TableFun/GraphFun.DUPS by TableFun/GraphFun.DUP;
wenzelm
parents:
23631
diff
changeset
|
419 |
error ("More than one " ^ name ^ " for " ^ quote c); |
237
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
175
diff
changeset
|
420 |
|
42253 | 421 |
fun lookup_tr tab c = Option.map fst (Symtab.lookup tab c); |
422 |
||
42288
2074b31650e6
discontinued special treatment of structure Syntax_Ext (formerly Syn_Ext);
wenzelm
parents:
42287
diff
changeset
|
423 |
fun remove_trtab trfuns = fold (Symtab.remove Syntax_Ext.eq_trfun) trfuns; |
21536
f119c730f509
extend_trtab: allow identical trfuns to be overwritten;
wenzelm
parents:
20784
diff
changeset
|
424 |
|
29004 | 425 |
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:
23631
diff
changeset
|
426 |
handle Symtab.DUP c => err_dup_trfun name c; |
237
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
175
diff
changeset
|
427 |
|
42288
2074b31650e6
discontinued special treatment of structure Syntax_Ext (formerly Syn_Ext);
wenzelm
parents:
42287
diff
changeset
|
428 |
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:
23631
diff
changeset
|
429 |
handle Symtab.DUP c => err_dup_trfun name c; |
237
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
175
diff
changeset
|
430 |
|
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
175
diff
changeset
|
431 |
|
5692 | 432 |
(* print (ast) translations *) |
433 |
||
42254 | 434 |
fun apply_tr' tab c ctxt T args = |
435 |
let |
|
436 |
val fns = map fst (Symtab.lookup_list tab c); |
|
437 |
fun app_first [] = raise Match |
|
438 |
| app_first (f :: fs) = f ctxt T args handle Match => app_first fs; |
|
439 |
in app_first fns end; |
|
440 |
||
441 |
fun apply_ast_tr' tab c ctxt args = |
|
442 |
let |
|
443 |
val fns = map fst (Symtab.lookup_list tab c); |
|
444 |
fun app_first [] = raise Match |
|
445 |
| app_first (f :: fs) = f ctxt args handle Match => app_first fs; |
|
446 |
in app_first fns end; |
|
42253 | 447 |
|
42288
2074b31650e6
discontinued special treatment of structure Syntax_Ext (formerly Syn_Ext);
wenzelm
parents:
42287
diff
changeset
|
448 |
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:
42287
diff
changeset
|
449 |
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:
42287
diff
changeset
|
450 |
fun merge_tr'tabs tab1 tab2 = Symtab.merge_list Syntax_Ext.eq_trfun (tab1, tab2); |
5692 | 451 |
|
452 |
||
237
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
175
diff
changeset
|
453 |
|
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
175
diff
changeset
|
454 |
(** tables of translation rules **) |
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
175
diff
changeset
|
455 |
|
5692 | 456 |
type ruletab = (Ast.ast * Ast.ast) list Symtab.table; |
237
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
175
diff
changeset
|
457 |
|
81005 | 458 |
fun dest_ruletab tab = Symtab.dest tab |> sort_by #1 |> maps #2; |
237
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
175
diff
changeset
|
459 |
|
81152 | 460 |
val update_ruletab = fold_rev (fn r => Symtab.update_list (op =) (Ast.rule_index r, r)); |
461 |
val remove_ruletab = fold (fn r => Symtab.remove_list (op =) (Ast.rule_index r, r)); |
|
18931 | 462 |
fun merge_ruletabs tab1 tab2 = Symtab.merge_list (op =) (tab1, tab2); |
0 | 463 |
|
464 |
||
465 |
||
466 |
(** datatype syntax **) |
|
467 |
||
237
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
175
diff
changeset
|
468 |
datatype syntax = |
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
175
diff
changeset
|
469 |
Syntax of { |
42288
2074b31650e6
discontinued special treatment of structure Syntax_Ext (formerly Syn_Ext);
wenzelm
parents:
42287
diff
changeset
|
470 |
input: Syntax_Ext.xprod list, |
81588
81a72b7fcb0c
minor performance tuning: avoid excessive Library.insert and Symbol.explode operations;
wenzelm
parents:
81176
diff
changeset
|
471 |
keywords: keywords, |
78009
f906f7f83dae
performance tuning: cached non-persistent Parser.gram reduces heap size by approx. 1-4%;
wenzelm
parents:
76082
diff
changeset
|
472 |
gram: Parser.gram Synchronized.cache, |
80748
43c4817375bf
support for syntax const dependencies, with minimal integrity checks;
wenzelm
parents:
80742
diff
changeset
|
473 |
consts: unit Graph.T, |
2913 | 474 |
prmodes: string list, |
21772 | 475 |
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:
175
diff
changeset
|
476 |
parse_ruletab: ruletab, |
21772 | 477 |
parse_trtab: ((Proof.context -> term list -> term) * stamp) Symtab.table, |
42247
12fe41a92cd5
typed_print_translation: discontinued show_sorts argument;
wenzelm
parents:
42245
diff
changeset
|
478 |
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:
175
diff
changeset
|
479 |
print_ruletab: ruletab, |
21772 | 480 |
print_ast_trtab: ((Proof.context -> Ast.ast list -> Ast.ast) * stamp) list Symtab.table, |
17079 | 481 |
prtabs: Printer.prtabs} * stamp; |
0 | 482 |
|
80073
40f5ddeda2b4
further performance tuning (after f906f7f83dae): interactive mode is closer to earlier approach with Lazy.value, which could be relevant with rather complex grammars under tight memory situations;
wenzelm
parents:
78013
diff
changeset
|
483 |
val cache_persistent = Unsynchronized.ref false; |
40f5ddeda2b4
further performance tuning (after f906f7f83dae): interactive mode is closer to earlier approach with Lazy.value, which could be relevant with rather complex grammars under tight memory situations;
wenzelm
parents:
78013
diff
changeset
|
484 |
|
40f5ddeda2b4
further performance tuning (after f906f7f83dae): interactive mode is closer to earlier approach with Lazy.value, which could be relevant with rather complex grammars under tight memory situations;
wenzelm
parents:
78013
diff
changeset
|
485 |
fun cache_eval (gram: Parser.gram Synchronized.cache) = |
40f5ddeda2b4
further performance tuning (after f906f7f83dae): interactive mode is closer to earlier approach with Lazy.value, which could be relevant with rather complex grammars under tight memory situations;
wenzelm
parents:
78013
diff
changeset
|
486 |
Synchronized.cache_eval {persistent = ! cache_persistent} gram; |
40f5ddeda2b4
further performance tuning (after f906f7f83dae): interactive mode is closer to earlier approach with Lazy.value, which could be relevant with rather complex grammars under tight memory situations;
wenzelm
parents:
78013
diff
changeset
|
487 |
|
40f5ddeda2b4
further performance tuning (after f906f7f83dae): interactive mode is closer to earlier approach with Lazy.value, which could be relevant with rather complex grammars under tight memory situations;
wenzelm
parents:
78013
diff
changeset
|
488 |
fun cache_syntax (Syntax ({gram, ...}, _)) = ignore (cache_eval gram); |
40f5ddeda2b4
further performance tuning (after f906f7f83dae): interactive mode is closer to earlier approach with Lazy.value, which could be relevant with rather complex grammars under tight memory situations;
wenzelm
parents:
78013
diff
changeset
|
489 |
|
17079 | 490 |
fun eq_syntax (Syntax (_, s1), Syntax (_, s2)) = s1 = s2; |
491 |
||
69077
11529ae45786
more approximative prefix syntax, including binder;
wenzelm
parents:
69071
diff
changeset
|
492 |
datatype approx = Prefix of string | Infix of {assoc: Printer.assoc, delim: string, pri: int}; |
11529ae45786
more approximative prefix syntax, including binder;
wenzelm
parents:
69071
diff
changeset
|
493 |
fun get_approx (Syntax ({prtabs, ...}, _)) c = |
11529ae45786
more approximative prefix syntax, including binder;
wenzelm
parents:
69071
diff
changeset
|
494 |
(case Printer.get_infix prtabs c of |
11529ae45786
more approximative prefix syntax, including binder;
wenzelm
parents:
69071
diff
changeset
|
495 |
SOME infx => SOME (Infix infx) |
11529ae45786
more approximative prefix syntax, including binder;
wenzelm
parents:
69071
diff
changeset
|
496 |
| NONE => |
11529ae45786
more approximative prefix syntax, including binder;
wenzelm
parents:
69071
diff
changeset
|
497 |
(case Printer.get_prefix prtabs c of |
11529ae45786
more approximative prefix syntax, including binder;
wenzelm
parents:
69071
diff
changeset
|
498 |
SOME prfx => SOME prfx |
69078
a5e904112ea9
more accurate syntax: e.g. avoid brackets as prefix notation;
wenzelm
parents:
69077
diff
changeset
|
499 |
| NONE => Printer.get_binder prtabs c) |
69077
11529ae45786
more approximative prefix syntax, including binder;
wenzelm
parents:
69071
diff
changeset
|
500 |
|> Option.map Prefix); |
69003 | 501 |
|
80767
079fe4292526
clarified Syntax.is_const (after 43c4817375bf): exclude logical consts from 'syntax_consts' / 'syntax_types', e.g. relevant for @{syntax_const} antiquotation;
wenzelm
parents:
80751
diff
changeset
|
502 |
fun is_const (Syntax ({consts, ...}, _)) c = |
81596 | 503 |
Graph.defined consts c andalso not (Lexicon.is_marked_entity c); |
80767
079fe4292526
clarified Syntax.is_const (after 43c4817375bf): exclude logical consts from 'syntax_consts' / 'syntax_types', e.g. relevant for @{syntax_const} antiquotation;
wenzelm
parents:
80751
diff
changeset
|
504 |
|
81588
81a72b7fcb0c
minor performance tuning: avoid excessive Library.insert and Symbol.explode operations;
wenzelm
parents:
81176
diff
changeset
|
505 |
fun is_keyword (Syntax ({keywords, ...}, _)) = member_keywords keywords; |
81a72b7fcb0c
minor performance tuning: avoid excessive Library.insert and Symbol.explode operations;
wenzelm
parents:
81176
diff
changeset
|
506 |
fun tokenize (Syntax ({keywords, ...}, _)) = tokenize_keywords keywords; |
80073
40f5ddeda2b4
further performance tuning (after f906f7f83dae): interactive mode is closer to earlier approach with Lazy.value, which could be relevant with rather complex grammars under tight memory situations;
wenzelm
parents:
78013
diff
changeset
|
507 |
fun parse (Syntax ({gram, ...}, _)) = Parser.parse (cache_eval gram); |
14687 | 508 |
|
42253 | 509 |
fun parse_ast_translation (Syntax ({parse_ast_trtab, ...}, _)) = lookup_tr parse_ast_trtab; |
510 |
fun parse_translation (Syntax ({parse_trtab, ...}, _)) = lookup_tr parse_trtab; |
|
42255 | 511 |
fun parse_rules (Syntax ({parse_ruletab, ...}, _)) = Symtab.lookup_list parse_ruletab; |
42254 | 512 |
fun print_translation (Syntax ({print_trtab, ...}, _)) = apply_tr' print_trtab; |
42255 | 513 |
fun print_rules (Syntax ({print_ruletab, ...}, _)) = Symtab.lookup_list print_ruletab; |
42254 | 514 |
fun print_ast_translation (Syntax ({print_ast_trtab, ...}, _)) = apply_ast_tr' print_ast_trtab; |
42253 | 515 |
|
42255 | 516 |
fun prtabs (Syntax ({prtabs, ...}, _)) = prtabs; |
81176
c0522b2d3df6
clarified signature: more explicit Syntax.print_mode_tabs, depending on print_mode_value ();
wenzelm
parents:
81163
diff
changeset
|
517 |
val print_mode_tabs = Printer.print_mode_tabs o prtabs; |
42255 | 518 |
|
20784 | 519 |
type mode = string * bool; |
24970 | 520 |
val mode_default = ("", true); |
37146
f652333bbf8e
renamed structure PrintMode to Print_Mode, keeping the old name as legacy alias for some time;
wenzelm
parents:
36739
diff
changeset
|
521 |
val mode_input = (Print_Mode.input, true); |
15755
50ac97ebe7d8
expect translations functions to be stamped already;
wenzelm
parents:
15574
diff
changeset
|
522 |
|
78009
f906f7f83dae
performance tuning: cached non-persistent Parser.gram reduces heap size by approx. 1-4%;
wenzelm
parents:
76082
diff
changeset
|
523 |
fun extend_gram new_xprods old_xprods gram = |
f906f7f83dae
performance tuning: cached non-persistent Parser.gram reduces heap size by approx. 1-4%;
wenzelm
parents:
76082
diff
changeset
|
524 |
Synchronized.cache (fn () => |
f906f7f83dae
performance tuning: cached non-persistent Parser.gram reduces heap size by approx. 1-4%;
wenzelm
parents:
76082
diff
changeset
|
525 |
Parser.make_gram new_xprods old_xprods (Synchronized.cache_peek gram)); |
f906f7f83dae
performance tuning: cached non-persistent Parser.gram reduces heap size by approx. 1-4%;
wenzelm
parents:
76082
diff
changeset
|
526 |
|
f906f7f83dae
performance tuning: cached non-persistent Parser.gram reduces heap size by approx. 1-4%;
wenzelm
parents:
76082
diff
changeset
|
527 |
fun new_gram new_xprods = |
f906f7f83dae
performance tuning: cached non-persistent Parser.gram reduces heap size by approx. 1-4%;
wenzelm
parents:
76082
diff
changeset
|
528 |
Synchronized.cache (fn () => Parser.make_gram new_xprods [] NONE); |
f906f7f83dae
performance tuning: cached non-persistent Parser.gram reduces heap size by approx. 1-4%;
wenzelm
parents:
76082
diff
changeset
|
529 |
|
18 | 530 |
|
80742 | 531 |
(* syntax consts *) |
532 |
||
80748
43c4817375bf
support for syntax const dependencies, with minimal integrity checks;
wenzelm
parents:
80742
diff
changeset
|
533 |
fun err_cyclic_consts css = |
43c4817375bf
support for syntax const dependencies, with minimal integrity checks;
wenzelm
parents:
80742
diff
changeset
|
534 |
error (cat_lines (map (fn cs => |
43c4817375bf
support for syntax const dependencies, with minimal integrity checks;
wenzelm
parents:
80742
diff
changeset
|
535 |
"Cycle in syntax consts: " ^ (space_implode " \<leadsto> " (map quote cs))) css)); |
43c4817375bf
support for syntax const dependencies, with minimal integrity checks;
wenzelm
parents:
80742
diff
changeset
|
536 |
|
43c4817375bf
support for syntax const dependencies, with minimal integrity checks;
wenzelm
parents:
80742
diff
changeset
|
537 |
fun get_consts (Syntax ({consts, ...}, _)) c = |
81163 | 538 |
if Graph.defined consts c |
539 |
then filter (Graph.is_minimal consts) (Graph.all_preds consts [c]) |
|
80751
c7f7e58509af
more accurate markup (amending 43c4817375bf): only consider primitive syntax consts, avoid extra Markup.intensify e.g. due to "\<^const>Pure.all_binder";
wenzelm
parents:
80748
diff
changeset
|
540 |
else [c]; |
80742 | 541 |
|
81592 | 542 |
fun add_consts (c, bs) consts = |
81596 | 543 |
if c = "" orelse (null bs andalso (Lexicon.is_marked_entity c orelse Graph.defined consts c)) |
80748
43c4817375bf
support for syntax const dependencies, with minimal integrity checks;
wenzelm
parents:
80742
diff
changeset
|
544 |
then consts |
43c4817375bf
support for syntax const dependencies, with minimal integrity checks;
wenzelm
parents:
80742
diff
changeset
|
545 |
else |
43c4817375bf
support for syntax const dependencies, with minimal integrity checks;
wenzelm
parents:
80742
diff
changeset
|
546 |
consts |
43c4817375bf
support for syntax const dependencies, with minimal integrity checks;
wenzelm
parents:
80742
diff
changeset
|
547 |
|> fold (fn a => Graph.default_node (a, ())) (c :: bs) |
43c4817375bf
support for syntax const dependencies, with minimal integrity checks;
wenzelm
parents:
80742
diff
changeset
|
548 |
|> Graph.add_deps_acyclic (c, bs) |
43c4817375bf
support for syntax const dependencies, with minimal integrity checks;
wenzelm
parents:
80742
diff
changeset
|
549 |
handle Graph.CYCLES css => err_cyclic_consts css; |
80742 | 550 |
|
551 |
||
237
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
175
diff
changeset
|
552 |
(* empty_syntax *) |
18 | 553 |
|
17079 | 554 |
val empty_syntax = Syntax |
555 |
({input = [], |
|
81588
81a72b7fcb0c
minor performance tuning: avoid excessive Library.insert and Symbol.explode operations;
wenzelm
parents:
81176
diff
changeset
|
556 |
keywords = empty_keywords, |
78009
f906f7f83dae
performance tuning: cached non-persistent Parser.gram reduces heap size by approx. 1-4%;
wenzelm
parents:
76082
diff
changeset
|
557 |
gram = Synchronized.cache (fn () => Parser.empty_gram), |
80748
43c4817375bf
support for syntax const dependencies, with minimal integrity checks;
wenzelm
parents:
80742
diff
changeset
|
558 |
consts = Graph.empty, |
2913 | 559 |
prmodes = [], |
5692 | 560 |
parse_ast_trtab = Symtab.empty, |
561 |
parse_ruletab = Symtab.empty, |
|
562 |
parse_trtab = Symtab.empty, |
|
563 |
print_trtab = Symtab.empty, |
|
564 |
print_ruletab = Symtab.empty, |
|
565 |
print_ast_trtab = Symtab.empty, |
|
17079 | 566 |
prtabs = Printer.empty_prtabs}, stamp ()); |
167 | 567 |
|
568 |
||
25394
db25c98f32e1
syntax operations: turned extend'' into update'' (absorb duplicates);
wenzelm
parents:
25387
diff
changeset
|
569 |
(* update_syntax *) |
167 | 570 |
|
25394
db25c98f32e1
syntax operations: turned extend'' into update'' (absorb duplicates);
wenzelm
parents:
25387
diff
changeset
|
571 |
fun update_syntax (mode, inout) syn_ext (Syntax (tabs, _)) = |
167 | 572 |
let |
81588
81a72b7fcb0c
minor performance tuning: avoid excessive Library.insert and Symbol.explode operations;
wenzelm
parents:
81176
diff
changeset
|
573 |
val {input, keywords, gram, consts = consts1, prmodes, parse_ast_trtab, parse_ruletab, |
42268 | 574 |
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:
42287
diff
changeset
|
575 |
val Syntax_Ext.Syn_Ext {xprods, consts = consts2, parse_ast_translation, |
42268 | 576 |
parse_rules, parse_translation, print_translation, print_rules, |
577 |
print_ast_translation} = syn_ext; |
|
36208
74c5e6e3c1d3
update_syntax: add new productions only once, to allow repeated local notation, for example;
wenzelm
parents:
35668
diff
changeset
|
578 |
val new_xprods = |
74c5e6e3c1d3
update_syntax: add new productions only once, to allow repeated local notation, for example;
wenzelm
parents:
35668
diff
changeset
|
579 |
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:
19482
diff
changeset
|
580 |
fun if_inout xs = if inout then xs else []; |
167 | 581 |
in |
17079 | 582 |
Syntax |
36208
74c5e6e3c1d3
update_syntax: add new productions only once, to allow repeated local notation, for example;
wenzelm
parents:
35668
diff
changeset
|
583 |
({input = new_xprods @ input, |
81588
81a72b7fcb0c
minor performance tuning: avoid excessive Library.insert and Symbol.explode operations;
wenzelm
parents:
81176
diff
changeset
|
584 |
keywords = (fold o Syntax_Ext.fold_delims) add_keywords new_xprods keywords, |
78009
f906f7f83dae
performance tuning: cached non-persistent Parser.gram reduces heap size by approx. 1-4%;
wenzelm
parents:
76082
diff
changeset
|
585 |
gram = if null new_xprods then gram else extend_gram new_xprods input gram, |
81592 | 586 |
consts = fold add_consts consts2 consts1, |
42268 | 587 |
prmodes = insert (op =) mode prmodes, |
167 | 588 |
parse_ast_trtab = |
25394
db25c98f32e1
syntax operations: turned extend'' into update'' (absorb duplicates);
wenzelm
parents:
25387
diff
changeset
|
589 |
update_trtab "parse ast translation" (if_inout parse_ast_translation) parse_ast_trtab, |
db25c98f32e1
syntax operations: turned extend'' into update'' (absorb duplicates);
wenzelm
parents:
25387
diff
changeset
|
590 |
parse_ruletab = update_ruletab (if_inout parse_rules) parse_ruletab, |
db25c98f32e1
syntax operations: turned extend'' into update'' (absorb duplicates);
wenzelm
parents:
25387
diff
changeset
|
591 |
parse_trtab = update_trtab "parse translation" (if_inout parse_translation) parse_trtab, |
db25c98f32e1
syntax operations: turned extend'' into update'' (absorb duplicates);
wenzelm
parents:
25387
diff
changeset
|
592 |
print_trtab = update_tr'tab print_translation print_trtab, |
db25c98f32e1
syntax operations: turned extend'' into update'' (absorb duplicates);
wenzelm
parents:
25387
diff
changeset
|
593 |
print_ruletab = update_ruletab print_rules print_ruletab, |
db25c98f32e1
syntax operations: turned extend'' into update'' (absorb duplicates);
wenzelm
parents:
25387
diff
changeset
|
594 |
print_ast_trtab = update_tr'tab print_ast_translation print_ast_trtab, |
db25c98f32e1
syntax operations: turned extend'' into update'' (absorb duplicates);
wenzelm
parents:
25387
diff
changeset
|
595 |
prtabs = Printer.update_prtabs mode xprods prtabs}, stamp ()) |
18 | 596 |
end; |
597 |
||
15755
50ac97ebe7d8
expect translations functions to be stamped already;
wenzelm
parents:
15574
diff
changeset
|
598 |
|
50ac97ebe7d8
expect translations functions to be stamped already;
wenzelm
parents:
15574
diff
changeset
|
599 |
(* remove_syntax *) |
50ac97ebe7d8
expect translations functions to be stamped already;
wenzelm
parents:
15574
diff
changeset
|
600 |
|
17079 | 601 |
fun remove_syntax (mode, inout) syn_ext (Syntax (tabs, _)) = |
15755
50ac97ebe7d8
expect translations functions to be stamped already;
wenzelm
parents:
15574
diff
changeset
|
602 |
let |
42288
2074b31650e6
discontinued special treatment of structure Syntax_Ext (formerly Syn_Ext);
wenzelm
parents:
42287
diff
changeset
|
603 |
val Syntax_Ext.Syn_Ext {xprods, consts = _, parse_ast_translation, parse_rules, |
42268 | 604 |
parse_translation, print_translation, print_rules, print_ast_translation} = syn_ext; |
81588
81a72b7fcb0c
minor performance tuning: avoid excessive Library.insert and Symbol.explode operations;
wenzelm
parents:
81176
diff
changeset
|
605 |
val {input, keywords, gram, consts, prmodes, parse_ast_trtab, parse_ruletab, |
42268 | 606 |
parse_trtab, print_trtab, print_ruletab, print_ast_trtab, prtabs} = tabs; |
19300 | 607 |
val input' = if inout then subtract (op =) xprods input else input; |
25394
db25c98f32e1
syntax operations: turned extend'' into update'' (absorb duplicates);
wenzelm
parents:
25387
diff
changeset
|
608 |
val changed = length input <> length input'; |
19546
00d5c7c7ce07
extend/remove_syntax: observe inout flag for translations, too;
wenzelm
parents:
19482
diff
changeset
|
609 |
fun if_inout xs = if inout then xs else []; |
15755
50ac97ebe7d8
expect translations functions to be stamped already;
wenzelm
parents:
15574
diff
changeset
|
610 |
in |
17079 | 611 |
Syntax |
612 |
({input = input', |
|
81588
81a72b7fcb0c
minor performance tuning: avoid excessive Library.insert and Symbol.explode operations;
wenzelm
parents:
81176
diff
changeset
|
613 |
keywords = |
81a72b7fcb0c
minor performance tuning: avoid excessive Library.insert and Symbol.explode operations;
wenzelm
parents:
81176
diff
changeset
|
614 |
if changed |
81a72b7fcb0c
minor performance tuning: avoid excessive Library.insert and Symbol.explode operations;
wenzelm
parents:
81176
diff
changeset
|
615 |
then make_keywords (Symset.build (input' |> (fold o Syntax_Ext.fold_delims) Symset.insert)) |
81a72b7fcb0c
minor performance tuning: avoid excessive Library.insert and Symbol.explode operations;
wenzelm
parents:
81176
diff
changeset
|
616 |
else keywords, |
78009
f906f7f83dae
performance tuning: cached non-persistent Parser.gram reduces heap size by approx. 1-4%;
wenzelm
parents:
76082
diff
changeset
|
617 |
gram = if changed then new_gram input' else gram, |
15755
50ac97ebe7d8
expect translations functions to be stamped already;
wenzelm
parents:
15574
diff
changeset
|
618 |
consts = consts, |
50ac97ebe7d8
expect translations functions to be stamped already;
wenzelm
parents:
15574
diff
changeset
|
619 |
prmodes = prmodes, |
19546
00d5c7c7ce07
extend/remove_syntax: observe inout flag for translations, too;
wenzelm
parents:
19482
diff
changeset
|
620 |
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:
19482
diff
changeset
|
621 |
parse_ruletab = remove_ruletab (if_inout parse_rules) parse_ruletab, |
00d5c7c7ce07
extend/remove_syntax: observe inout flag for translations, too;
wenzelm
parents:
19482
diff
changeset
|
622 |
parse_trtab = remove_trtab (if_inout parse_translation) parse_trtab, |
15755
50ac97ebe7d8
expect translations functions to be stamped already;
wenzelm
parents:
15574
diff
changeset
|
623 |
print_trtab = remove_tr'tab print_translation print_trtab, |
50ac97ebe7d8
expect translations functions to be stamped already;
wenzelm
parents:
15574
diff
changeset
|
624 |
print_ruletab = remove_ruletab print_rules print_ruletab, |
50ac97ebe7d8
expect translations functions to be stamped already;
wenzelm
parents:
15574
diff
changeset
|
625 |
print_ast_trtab = remove_tr'tab print_ast_translation print_ast_trtab, |
17079 | 626 |
prtabs = Printer.remove_prtabs mode xprods prtabs}, stamp ()) |
15755
50ac97ebe7d8
expect translations functions to be stamped already;
wenzelm
parents:
15574
diff
changeset
|
627 |
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:
14798
diff
changeset
|
628 |
|
18 | 629 |
|
45632
b23c42b9f78a
prefer Parser.make_gram over Parser.merge_gram, to approximate n-ary merges on theory import;
wenzelm
parents:
45490
diff
changeset
|
630 |
(* merge_syntax *) |
0 | 631 |
|
45632
b23c42b9f78a
prefer Parser.make_gram over Parser.merge_gram, to approximate n-ary merges on theory import;
wenzelm
parents:
45490
diff
changeset
|
632 |
fun merge_syntax (Syntax (tabs1, _), Syntax (tabs2, _)) = |
0 | 633 |
let |
81588
81a72b7fcb0c
minor performance tuning: avoid excessive Library.insert and Symbol.explode operations;
wenzelm
parents:
81176
diff
changeset
|
634 |
val {input = input1, keywords = keywords1, gram = gram1, consts = consts1, |
42268 | 635 |
prmodes = prmodes1, parse_ast_trtab = parse_ast_trtab1, parse_ruletab = parse_ruletab1, |
636 |
parse_trtab = parse_trtab1, print_trtab = print_trtab1, print_ruletab = print_ruletab1, |
|
637 |
print_ast_trtab = print_ast_trtab1, prtabs = prtabs1} = tabs1; |
|
237
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
175
diff
changeset
|
638 |
|
81588
81a72b7fcb0c
minor performance tuning: avoid excessive Library.insert and Symbol.explode operations;
wenzelm
parents:
81176
diff
changeset
|
639 |
val {input = input2, keywords = keywords2, gram = gram2, consts = consts2, |
42268 | 640 |
prmodes = prmodes2, parse_ast_trtab = parse_ast_trtab2, parse_ruletab = parse_ruletab2, |
641 |
parse_trtab = parse_trtab2, print_trtab = print_trtab2, print_ruletab = print_ruletab2, |
|
642 |
print_ast_trtab = print_ast_trtab2, prtabs = prtabs2} = tabs2; |
|
45632
b23c42b9f78a
prefer Parser.make_gram over Parser.merge_gram, to approximate n-ary merges on theory import;
wenzelm
parents:
45490
diff
changeset
|
643 |
|
b23c42b9f78a
prefer Parser.make_gram over Parser.merge_gram, to approximate n-ary merges on theory import;
wenzelm
parents:
45490
diff
changeset
|
644 |
val (input', gram') = |
70526 | 645 |
if pointer_eq (input1, input2) then (input1, gram1) |
646 |
else |
|
647 |
(case subtract (op =) input1 input2 of |
|
648 |
[] => (input1, gram1) |
|
649 |
| new_xprods2 => |
|
650 |
if subset (op =) (input1, input2) then (input2, gram2) |
|
651 |
else |
|
652 |
let |
|
653 |
val input' = new_xprods2 @ input1; |
|
78013
f5b67198b019
backed out changeset 6c2494750a4e: it hardly makes a difference for heap size, but crashes arm64_32-darwin for unknown reasons;
wenzelm
parents:
78010
diff
changeset
|
654 |
val gram' = new_gram input'; |
70526 | 655 |
in (input', gram') end); |
0 | 656 |
in |
17079 | 657 |
Syntax |
45632
b23c42b9f78a
prefer Parser.make_gram over Parser.merge_gram, to approximate n-ary merges on theory import;
wenzelm
parents:
45490
diff
changeset
|
658 |
({input = input', |
81588
81a72b7fcb0c
minor performance tuning: avoid excessive Library.insert and Symbol.explode operations;
wenzelm
parents:
81176
diff
changeset
|
659 |
keywords = merge_keywords (keywords1, keywords2), |
45632
b23c42b9f78a
prefer Parser.make_gram over Parser.merge_gram, to approximate n-ary merges on theory import;
wenzelm
parents:
45490
diff
changeset
|
660 |
gram = gram', |
80748
43c4817375bf
support for syntax const dependencies, with minimal integrity checks;
wenzelm
parents:
80742
diff
changeset
|
661 |
consts = |
43c4817375bf
support for syntax const dependencies, with minimal integrity checks;
wenzelm
parents:
80742
diff
changeset
|
662 |
Graph.merge_acyclic (op =) (consts1, consts2) |
43c4817375bf
support for syntax const dependencies, with minimal integrity checks;
wenzelm
parents:
80742
diff
changeset
|
663 |
handle Graph.CYCLES css => err_cyclic_consts css, |
18921 | 664 |
prmodes = Library.merge (op =) (prmodes1, prmodes2), |
237
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
175
diff
changeset
|
665 |
parse_ast_trtab = |
15755
50ac97ebe7d8
expect translations functions to be stamped already;
wenzelm
parents:
15574
diff
changeset
|
666 |
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:
175
diff
changeset
|
667 |
parse_ruletab = merge_ruletabs parse_ruletab1 parse_ruletab2, |
15755
50ac97ebe7d8
expect translations functions to be stamped already;
wenzelm
parents:
15574
diff
changeset
|
668 |
parse_trtab = merge_trtabs "parse translation" parse_trtab1 parse_trtab2, |
5692 | 669 |
print_trtab = merge_tr'tabs print_trtab1 print_trtab2, |
237
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
175
diff
changeset
|
670 |
print_ruletab = merge_ruletabs print_ruletab1 print_ruletab2, |
5692 | 671 |
print_ast_trtab = merge_tr'tabs print_ast_trtab1 print_ast_trtab2, |
81176
c0522b2d3df6
clarified signature: more explicit Syntax.print_mode_tabs, depending on print_mode_value ();
wenzelm
parents:
81163
diff
changeset
|
672 |
prtabs = Printer.merge_prtabs (prtabs1, prtabs2)}, stamp ()) |
0 | 673 |
end; |
674 |
||
675 |
||
4887 | 676 |
|
15759 | 677 |
(** print syntax **) |
678 |
||
679 |
local |
|
0 | 680 |
|
260 | 681 |
fun pretty_strs_qs name strs = |
28840 | 682 |
Pretty.strs (name :: map quote (sort_strings strs)); |
0 | 683 |
|
17079 | 684 |
fun pretty_gram (Syntax (tabs, _)) = |
0 | 685 |
let |
81588
81a72b7fcb0c
minor performance tuning: avoid excessive Library.insert and Symbol.explode operations;
wenzelm
parents:
81176
diff
changeset
|
686 |
val {keywords, prmodes, gram, ...} = tabs; |
28375 | 687 |
val prmodes' = sort_strings (filter_out (fn s => s = "") prmodes); |
0 | 688 |
in |
67513 | 689 |
[Pretty.block (Pretty.breaks |
81588
81a72b7fcb0c
minor performance tuning: avoid excessive Library.insert and Symbol.explode operations;
wenzelm
parents:
81176
diff
changeset
|
690 |
(Pretty.str "lexicon:" :: map (Pretty.quote o Pretty.keyword1) (dest_keywords keywords))), |
80073
40f5ddeda2b4
further performance tuning (after f906f7f83dae): interactive mode is closer to earlier approach with Lazy.value, which could be relevant with rather complex grammars under tight memory situations;
wenzelm
parents:
78013
diff
changeset
|
691 |
Pretty.big_list "productions:" (Parser.pretty_gram (cache_eval gram)), |
8720 | 692 |
pretty_strs_qs "print modes:" prmodes'] |
0 | 693 |
end; |
694 |
||
17079 | 695 |
fun pretty_trans (Syntax (tabs, _)) = |
0 | 696 |
let |
42298
d622145603ee
more accurate markup for syntax consts, notably binders which point back to the original logical entity;
wenzelm
parents:
42297
diff
changeset
|
697 |
fun pretty_tab name tab = |
d622145603ee
more accurate markup for syntax consts, notably binders which point back to the original logical entity;
wenzelm
parents:
42297
diff
changeset
|
698 |
pretty_strs_qs name (sort_strings (Symtab.keys tab)); |
0 | 699 |
|
260 | 700 |
fun pretty_ruletab name tab = |
51580
64ef8260dc60
Pretty.item markup for improved readability of lists of items;
wenzelm
parents:
50201
diff
changeset
|
701 |
Pretty.big_list name (map (Pretty.item o single o Ast.pretty_rule) (dest_ruletab tab)); |
0 | 702 |
|
237
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
175
diff
changeset
|
703 |
val {consts, parse_ast_trtab, parse_ruletab, parse_trtab, print_trtab, |
42268 | 704 |
print_ruletab, print_ast_trtab, ...} = tabs; |
0 | 705 |
in |
80748
43c4817375bf
support for syntax const dependencies, with minimal integrity checks;
wenzelm
parents:
80742
diff
changeset
|
706 |
[pretty_strs_qs "consts:" (sort_strings (Graph.keys consts)), |
42298
d622145603ee
more accurate markup for syntax consts, notably binders which point back to the original logical entity;
wenzelm
parents:
42297
diff
changeset
|
707 |
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:
42297
diff
changeset
|
708 |
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:
42297
diff
changeset
|
709 |
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:
42297
diff
changeset
|
710 |
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:
42297
diff
changeset
|
711 |
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:
42297
diff
changeset
|
712 |
pretty_tab "print_ast_translation:" print_ast_trtab] |
0 | 713 |
end; |
714 |
||
15759 | 715 |
in |
0 | 716 |
|
82587
7415414bd9d8
more scalable: discontinue odd shortcuts from 6b3739fee456, which produce bulky strings internally;
wenzelm
parents:
81993
diff
changeset
|
717 |
fun print_gram syn = Pretty.writeln (Pretty.chunks (pretty_gram syn)); |
7415414bd9d8
more scalable: discontinue odd shortcuts from 6b3739fee456, which produce bulky strings internally;
wenzelm
parents:
81993
diff
changeset
|
718 |
fun print_trans syn = Pretty.writeln (Pretty.chunks (pretty_trans syn)); |
7415414bd9d8
more scalable: discontinue odd shortcuts from 6b3739fee456, which produce bulky strings internally;
wenzelm
parents:
81993
diff
changeset
|
719 |
fun print_syntax syn = Pretty.writeln (Pretty.chunks (pretty_gram syn @ pretty_trans syn)); |
0 | 720 |
|
15759 | 721 |
end; |
0 | 722 |
|
723 |
||
724 |
||
1158 | 725 |
(** prepare translation rules **) |
726 |
||
42204 | 727 |
(* rules *) |
728 |
||
729 |
fun map_trrule f (Parse_Rule (x, y)) = Parse_Rule (f x, f y) |
|
730 |
| map_trrule f (Print_Rule (x, y)) = Print_Rule (f x, f y) |
|
731 |
| map_trrule f (Parse_Print_Rule (x, y)) = Parse_Print_Rule (f x, f y); |
|
1158 | 732 |
|
42204 | 733 |
fun parse_rule (Parse_Rule pats) = SOME pats |
734 |
| parse_rule (Print_Rule _) = NONE |
|
735 |
| parse_rule (Parse_Print_Rule pats) = SOME pats; |
|
1158 | 736 |
|
42204 | 737 |
fun print_rule (Parse_Rule _) = NONE |
738 |
| print_rule (Print_Rule pats) = SOME (swap pats) |
|
739 |
| print_rule (Parse_Print_Rule pats) = SOME (swap pats); |
|
1158 | 740 |
|
741 |
||
42204 | 742 |
(* check_rules *) |
743 |
||
19262 | 744 |
local |
745 |
||
42048
afd11ca8e018
support for encoded positions (for id_position, longid_position) as pseudo type-constraints -- still inactive;
wenzelm
parents:
42044
diff
changeset
|
746 |
fun check_rule rule = |
5692 | 747 |
(case Ast.rule_error rule of |
15531 | 748 |
SOME msg => |
1158 | 749 |
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:
42044
diff
changeset
|
750 |
Pretty.string_of (Ast.pretty_rule rule)) |
15531 | 751 |
| NONE => rule); |
888
3a1de9454d13
improved read_xrules: patterns no longer read twice;
wenzelm
parents:
882
diff
changeset
|
752 |
|
42204 | 753 |
in |
754 |
||
755 |
fun check_rules rules = |
|
756 |
(map check_rule (map_filter parse_rule rules), |
|
757 |
map check_rule (map_filter print_rule rules)); |
|
758 |
||
759 |
end; |
|
760 |
||
761 |
||
18 | 762 |
|
19262 | 763 |
(** modify syntax **) |
383
fcea89074e4c
added incremental extension functions: extend_log_types, extend_type_gram,
wenzelm
parents:
330
diff
changeset
|
764 |
|
81116
0fb1e2dd4122
first-class support for "unbundle no foobar_syntax" -- avoid redundant "bundle no_foobar_syntax" definitions;
wenzelm
parents:
81005
diff
changeset
|
765 |
(* updates *) |
0fb1e2dd4122
first-class support for "unbundle no foobar_syntax" -- avoid redundant "bundle no_foobar_syntax" definitions;
wenzelm
parents:
81005
diff
changeset
|
766 |
|
80897
5328d67ec647
more explicit context for syn_ext/mixfix operations, but it often degenerates to background theory;
wenzelm
parents:
80767
diff
changeset
|
767 |
fun update_trfuns ctxt = update_syntax mode_default o Syntax_Ext.syn_ext_trfuns ctxt; |
5692 | 768 |
|
80897
5328d67ec647
more explicit context for syn_ext/mixfix operations, but it often degenerates to background theory;
wenzelm
parents:
80767
diff
changeset
|
769 |
fun update_type_gram ctxt add prmode decls = |
5328d67ec647
more explicit context for syn_ext/mixfix operations, but it often degenerates to background theory;
wenzelm
parents:
80767
diff
changeset
|
770 |
(if add then update_syntax else remove_syntax) prmode (Mixfix.syn_ext_types ctxt decls); |
25387 | 771 |
|
80897
5328d67ec647
more explicit context for syn_ext/mixfix operations, but it often degenerates to background theory;
wenzelm
parents:
80767
diff
changeset
|
772 |
fun update_const_gram ctxt add logical_types prmode decls = |
5328d67ec647
more explicit context for syn_ext/mixfix operations, but it often degenerates to background theory;
wenzelm
parents:
80767
diff
changeset
|
773 |
(if add then update_syntax else remove_syntax) prmode |
5328d67ec647
more explicit context for syn_ext/mixfix operations, but it often degenerates to background theory;
wenzelm
parents:
80767
diff
changeset
|
774 |
(Mixfix.syn_ext_consts ctxt logical_types decls); |
15755
50ac97ebe7d8
expect translations functions to be stamped already;
wenzelm
parents:
15574
diff
changeset
|
775 |
|
80897
5328d67ec647
more explicit context for syn_ext/mixfix operations, but it often degenerates to background theory;
wenzelm
parents:
80767
diff
changeset
|
776 |
val update_consts = update_syntax mode_default oo Syntax_Ext.syn_ext_consts; |
80748
43c4817375bf
support for syntax const dependencies, with minimal integrity checks;
wenzelm
parents:
80742
diff
changeset
|
777 |
|
81590 | 778 |
fun update_trrules ctxt add = |
779 |
(if add then update_syntax else remove_syntax) mode_default o |
|
780 |
Syntax_Ext.syn_ext_rules ctxt o check_rules; |
|
5692 | 781 |
|
42476 | 782 |
|
81116
0fb1e2dd4122
first-class support for "unbundle no foobar_syntax" -- avoid redundant "bundle no_foobar_syntax" definitions;
wenzelm
parents:
81005
diff
changeset
|
783 |
(* polarity of declarations: true = add, false = del *) |
0fb1e2dd4122
first-class support for "unbundle no foobar_syntax" -- avoid redundant "bundle no_foobar_syntax" definitions;
wenzelm
parents:
81005
diff
changeset
|
784 |
|
0fb1e2dd4122
first-class support for "unbundle no foobar_syntax" -- avoid redundant "bundle no_foobar_syntax" definitions;
wenzelm
parents:
81005
diff
changeset
|
785 |
val polarity = Config.declare_bool ("polarity", Position.none) (K true); |
0fb1e2dd4122
first-class support for "unbundle no foobar_syntax" -- avoid redundant "bundle no_foobar_syntax" definitions;
wenzelm
parents:
81005
diff
changeset
|
786 |
|
0fb1e2dd4122
first-class support for "unbundle no foobar_syntax" -- avoid redundant "bundle no_foobar_syntax" definitions;
wenzelm
parents:
81005
diff
changeset
|
787 |
fun get_polarity ctxt = Config.get ctxt polarity; |
0fb1e2dd4122
first-class support for "unbundle no foobar_syntax" -- avoid redundant "bundle no_foobar_syntax" definitions;
wenzelm
parents:
81005
diff
changeset
|
788 |
val set_polarity = Config.put polarity; |
0fb1e2dd4122
first-class support for "unbundle no foobar_syntax" -- avoid redundant "bundle no_foobar_syntax" definitions;
wenzelm
parents:
81005
diff
changeset
|
789 |
val set_polarity_generic = Config.put_generic polarity; |
0fb1e2dd4122
first-class support for "unbundle no foobar_syntax" -- avoid redundant "bundle no_foobar_syntax" definitions;
wenzelm
parents:
81005
diff
changeset
|
790 |
|
0fb1e2dd4122
first-class support for "unbundle no foobar_syntax" -- avoid redundant "bundle no_foobar_syntax" definitions;
wenzelm
parents:
81005
diff
changeset
|
791 |
fun effective_polarity ctxt add = get_polarity ctxt = add; |
81993 | 792 |
fun effective_polarity_global thy add = Config.get_global thy polarity = add; |
793 |
val effective_polarity_generic = Context.cases effective_polarity_global effective_polarity; |
|
81116
0fb1e2dd4122
first-class support for "unbundle no foobar_syntax" -- avoid redundant "bundle no_foobar_syntax" definitions;
wenzelm
parents:
81005
diff
changeset
|
794 |
|
0fb1e2dd4122
first-class support for "unbundle no foobar_syntax" -- avoid redundant "bundle no_foobar_syntax" definitions;
wenzelm
parents:
81005
diff
changeset
|
795 |
|
42476 | 796 |
open Lexicon.Syntax; |
797 |
||
0 | 798 |
end; |