author | wenzelm |
Mon, 20 Sep 2010 15:08:51 +0200 | |
changeset 39555 | ccb223a4d49c |
parent 39510 | d9f5f01faa1b |
child 40959 | 49765c1104d4 |
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 |
||
260 | 8 |
signature BASIC_SYNTAX = |
2383 | 9 |
sig |
260 | 10 |
include AST0 |
556 | 11 |
include SYN_TRANS0 |
12 |
include MIXFIX0 |
|
260 | 13 |
include PRINTER0 |
2383 | 14 |
end; |
260 | 15 |
|
0 | 16 |
signature SYNTAX = |
2383 | 17 |
sig |
260 | 18 |
include AST1 |
0 | 19 |
include LEXICON0 |
237
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
175
diff
changeset
|
20 |
include SYN_EXT0 |
0 | 21 |
include TYPE_EXT0 |
556 | 22 |
include SYN_TRANS1 |
23 |
include MIXFIX1 |
|
0 | 24 |
include PRINTER0 |
30573 | 25 |
val read_token: string -> Symbol_Pos.T list * Position.T |
38238 | 26 |
val parse_token: Proof.context -> Markup.T -> string -> Symbol_Pos.T list * Position.T |
24263 | 27 |
val parse_sort: Proof.context -> string -> sort |
28 |
val parse_typ: Proof.context -> string -> typ |
|
29 |
val parse_term: Proof.context -> string -> term |
|
30 |
val parse_prop: Proof.context -> string -> term |
|
24768 | 31 |
val unparse_sort: Proof.context -> sort -> Pretty.T |
24923 | 32 |
val unparse_classrel: Proof.context -> class list -> Pretty.T |
33 |
val unparse_arity: Proof.context -> arity -> Pretty.T |
|
24768 | 34 |
val unparse_typ: Proof.context -> typ -> Pretty.T |
35 |
val unparse_term: Proof.context -> term -> Pretty.T |
|
36 |
val install_operations: |
|
37 |
{parse_sort: Proof.context -> string -> sort, |
|
38 |
parse_typ: Proof.context -> string -> typ, |
|
39 |
parse_term: Proof.context -> string -> term, |
|
40 |
parse_prop: Proof.context -> string -> term, |
|
41 |
unparse_sort: Proof.context -> sort -> Pretty.T, |
|
42 |
unparse_typ: Proof.context -> typ -> Pretty.T, |
|
43 |
unparse_term: Proof.context -> term -> Pretty.T} -> unit |
|
44 |
val print_checks: Proof.context -> unit |
|
45 |
val add_typ_check: int -> string -> |
|
25060 | 46 |
(typ list -> Proof.context -> (typ list * Proof.context) option) -> |
24768 | 47 |
Context.generic -> Context.generic |
48 |
val add_term_check: int -> string -> |
|
25060 | 49 |
(term list -> Proof.context -> (term list * Proof.context) option) -> |
24768 | 50 |
Context.generic -> Context.generic |
51 |
val add_typ_uncheck: int -> string -> |
|
25060 | 52 |
(typ list -> Proof.context -> (typ list * Proof.context) option) -> |
24768 | 53 |
Context.generic -> Context.generic |
54 |
val add_term_uncheck: int -> string -> |
|
25060 | 55 |
(term list -> Proof.context -> (term list * Proof.context) option) -> |
24768 | 56 |
Context.generic -> Context.generic |
24488 | 57 |
val check_sort: Proof.context -> sort -> sort |
24512 | 58 |
val check_typ: Proof.context -> typ -> typ |
59 |
val check_term: Proof.context -> term -> term |
|
60 |
val check_prop: Proof.context -> term -> term |
|
24488 | 61 |
val check_typs: Proof.context -> typ list -> typ list |
24263 | 62 |
val check_terms: Proof.context -> term list -> term list |
63 |
val check_props: Proof.context -> term list -> term list |
|
24923 | 64 |
val uncheck_sort: Proof.context -> sort -> sort |
65 |
val uncheck_arity: Proof.context -> arity -> arity |
|
66 |
val uncheck_classrel: Proof.context -> class list -> class list |
|
24768 | 67 |
val uncheck_typs: Proof.context -> typ list -> typ list |
68 |
val uncheck_terms: Proof.context -> term list -> term list |
|
24263 | 69 |
val read_sort: Proof.context -> string -> sort |
70 |
val read_typ: Proof.context -> string -> typ |
|
24488 | 71 |
val read_term: Proof.context -> string -> term |
72 |
val read_prop: Proof.context -> string -> term |
|
24263 | 73 |
val read_terms: Proof.context -> string list -> term list |
74 |
val read_props: Proof.context -> string list -> term list |
|
24709 | 75 |
val read_sort_global: theory -> string -> sort |
76 |
val read_typ_global: theory -> string -> typ |
|
77 |
val read_term_global: theory -> string -> term |
|
78 |
val read_prop_global: theory -> string -> term |
|
24923 | 79 |
val pretty_term: Proof.context -> term -> Pretty.T |
80 |
val pretty_typ: Proof.context -> typ -> Pretty.T |
|
81 |
val pretty_sort: Proof.context -> sort -> Pretty.T |
|
82 |
val pretty_classrel: Proof.context -> class list -> Pretty.T |
|
83 |
val pretty_arity: Proof.context -> arity -> Pretty.T |
|
84 |
val string_of_term: Proof.context -> term -> string |
|
85 |
val string_of_typ: Proof.context -> typ -> string |
|
86 |
val string_of_sort: Proof.context -> sort -> string |
|
87 |
val string_of_classrel: Proof.context -> class list -> string |
|
88 |
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
|
89 |
val is_pretty_global: Proof.context -> bool |
030e4a818b39
moved global pretty/string_of functions from Sign to Syntax;
wenzelm
parents:
26704
diff
changeset
|
90 |
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
|
91 |
val init_pretty_global: theory -> Proof.context |
030e4a818b39
moved global pretty/string_of functions from Sign to Syntax;
wenzelm
parents:
26704
diff
changeset
|
92 |
val pretty_term_global: theory -> term -> Pretty.T |
030e4a818b39
moved global pretty/string_of functions from Sign to Syntax;
wenzelm
parents:
26704
diff
changeset
|
93 |
val pretty_typ_global: theory -> typ -> Pretty.T |
030e4a818b39
moved global pretty/string_of functions from Sign to Syntax;
wenzelm
parents:
26704
diff
changeset
|
94 |
val pretty_sort_global: theory -> sort -> Pretty.T |
030e4a818b39
moved global pretty/string_of functions from Sign to Syntax;
wenzelm
parents:
26704
diff
changeset
|
95 |
val string_of_term_global: theory -> term -> string |
030e4a818b39
moved global pretty/string_of functions from Sign to Syntax;
wenzelm
parents:
26704
diff
changeset
|
96 |
val string_of_typ_global: theory -> typ -> string |
030e4a818b39
moved global pretty/string_of functions from Sign to Syntax;
wenzelm
parents:
26704
diff
changeset
|
97 |
val string_of_sort_global: theory -> sort -> string |
24970 | 98 |
val pp: Proof.context -> Pretty.pp |
26951
030e4a818b39
moved global pretty/string_of functions from Sign to Syntax;
wenzelm
parents:
26704
diff
changeset
|
99 |
val pp_global: theory -> Pretty.pp |
39135
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
100 |
type syntax |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
101 |
val eq_syntax: syntax * syntax -> bool |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
102 |
val is_keyword: syntax -> string -> bool |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
103 |
type mode |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
104 |
val mode_default: mode |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
105 |
val mode_input: mode |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
106 |
val merge_syntaxes: syntax -> syntax -> syntax |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
107 |
val basic_syntax: syntax |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
108 |
val basic_nonterms: string list |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
109 |
val print_gram: syntax -> unit |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
110 |
val print_trans: syntax -> unit |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
111 |
val print_syntax: syntax -> unit |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
112 |
val guess_infix: syntax -> string -> mixfix option |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
113 |
val ambiguity_enabled: bool Config.T |
39163
4d701c0388c3
more explicit indication of Config.raw options, which are only needed for bootstrapping Pure;
wenzelm
parents:
39137
diff
changeset
|
114 |
val ambiguity_level_raw: Config.raw |
39135
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
115 |
val ambiguity_level: int Config.T |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
116 |
val ambiguity_limit: int Config.T |
39136
b0b3b6318e3b
Syntax.standard_parse_term: eliminated redundant Pretty.pp;
wenzelm
parents:
39135
diff
changeset
|
117 |
val standard_parse_term: (term -> string option) -> |
39135
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
118 |
(((string * int) * sort) list -> string * int -> Term.sort) -> |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
119 |
(string -> bool * string) -> (string -> string option) -> Proof.context -> |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
120 |
(string -> bool) -> syntax -> typ -> Symbol_Pos.T list * Position.T -> term |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
121 |
val standard_parse_typ: Proof.context -> syntax -> |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
122 |
((indexname * sort) list -> indexname -> sort) -> Symbol_Pos.T list * Position.T -> typ |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
123 |
val standard_parse_sort: Proof.context -> syntax -> Symbol_Pos.T list * Position.T -> sort |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
124 |
datatype 'a trrule = |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
125 |
ParseRule of 'a * 'a | |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
126 |
PrintRule of 'a * 'a | |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
127 |
ParsePrintRule of 'a * 'a |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
128 |
val map_trrule: ('a -> 'b) -> 'a trrule -> 'b trrule |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
129 |
val is_const: syntax -> string -> bool |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
130 |
val standard_unparse_term: {structs: string list, fixes: string list} -> |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
131 |
{extern_class: string -> xstring, extern_type: string -> xstring, |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
132 |
extern_const: string -> xstring} -> Proof.context -> syntax -> bool -> term -> Pretty.T |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
133 |
val standard_unparse_typ: {extern_class: string -> xstring, extern_type: string -> xstring} -> |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
134 |
Proof.context -> syntax -> typ -> Pretty.T |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
135 |
val standard_unparse_sort: {extern_class: string -> xstring} -> |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
136 |
Proof.context -> syntax -> sort -> Pretty.T |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
137 |
val update_trfuns: |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
138 |
(string * ((ast list -> ast) * stamp)) list * |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
139 |
(string * ((term list -> term) * stamp)) list * |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
140 |
(string * ((bool -> typ -> term list -> term) * stamp)) list * |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
141 |
(string * ((ast list -> ast) * stamp)) list -> syntax -> syntax |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
142 |
val update_advanced_trfuns: |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
143 |
(string * ((Proof.context -> ast list -> ast) * stamp)) list * |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
144 |
(string * ((Proof.context -> term list -> term) * stamp)) list * |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
145 |
(string * ((Proof.context -> bool -> typ -> term list -> term) * stamp)) list * |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
146 |
(string * ((Proof.context -> ast list -> ast) * stamp)) list -> syntax -> syntax |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
147 |
val extend_tokentrfuns: (string * string * (Proof.context -> string -> Pretty.T)) list -> |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
148 |
syntax -> syntax |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
149 |
val update_type_gram: bool -> mode -> (string * typ * mixfix) list -> syntax -> syntax |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
150 |
val update_const_gram: bool -> (string -> bool) -> |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
151 |
mode -> (string * typ * mixfix) list -> syntax -> syntax |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
152 |
val update_trrules: Proof.context -> (string -> bool) -> syntax -> |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
153 |
(string * string) trrule list -> syntax -> syntax |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
154 |
val remove_trrules: Proof.context -> (string -> bool) -> syntax -> |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
155 |
(string * string) trrule list -> syntax -> syntax |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
156 |
val update_trrules_i: ast trrule list -> syntax -> syntax |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
157 |
val remove_trrules_i: ast trrule list -> syntax -> syntax |
2383 | 158 |
end; |
0 | 159 |
|
12094 | 160 |
structure Syntax: SYNTAX = |
0 | 161 |
struct |
162 |
||
39135
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
163 |
(** inner syntax operations **) |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
164 |
|
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
165 |
(* read token -- with optional YXML encoding of position *) |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
166 |
|
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
167 |
fun read_token str = |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
168 |
let |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
169 |
val tree = YXML.parse str handle Fail msg => error msg; |
39555
ccb223a4d49c
added XML.content_of convenience -- cover XML.body, which is the general situation;
wenzelm
parents:
39510
diff
changeset
|
170 |
val text = XML.content_of [tree]; |
39135
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
171 |
val pos = |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
172 |
(case tree of |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
173 |
XML.Elem ((name, props), _) => |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
174 |
if name = Markup.tokenN then Position.of_properties props |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
175 |
else Position.none |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
176 |
| XML.Text _ => Position.none); |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
177 |
in (Symbol_Pos.explode (text, pos), pos) end; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
178 |
|
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
179 |
|
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
180 |
(* (un)parsing *) |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
181 |
|
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
182 |
fun parse_token ctxt markup str = |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
183 |
let |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
184 |
val (syms, pos) = read_token str; |
39507
839873937ddd
tuned signature of (Context_)Position.report variants;
wenzelm
parents:
39288
diff
changeset
|
185 |
val _ = Context_Position.report ctxt pos markup; |
39135
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
186 |
in (syms, pos) end; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
187 |
|
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
188 |
local |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
189 |
|
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
190 |
type operations = |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
191 |
{parse_sort: Proof.context -> string -> sort, |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
192 |
parse_typ: Proof.context -> string -> typ, |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
193 |
parse_term: Proof.context -> string -> term, |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
194 |
parse_prop: Proof.context -> string -> term, |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
195 |
unparse_sort: Proof.context -> sort -> Pretty.T, |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
196 |
unparse_typ: Proof.context -> typ -> Pretty.T, |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
197 |
unparse_term: Proof.context -> term -> Pretty.T}; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
198 |
|
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
199 |
val operations: operations Single_Assignment.var = Single_Assignment.var "Syntax.operations"; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
200 |
|
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
201 |
fun operation which ctxt x = |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
202 |
(case Single_Assignment.peek operations of |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
203 |
NONE => raise Fail "Inner syntax operations not installed" |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
204 |
| SOME ops => which ops ctxt x); |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
205 |
|
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
206 |
in |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
207 |
|
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
208 |
val parse_sort = operation #parse_sort; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
209 |
val parse_typ = operation #parse_typ; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
210 |
val parse_term = operation #parse_term; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
211 |
val parse_prop = operation #parse_prop; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
212 |
val unparse_sort = operation #unparse_sort; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
213 |
val unparse_typ = operation #unparse_typ; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
214 |
val unparse_term = operation #unparse_term; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
215 |
|
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
216 |
fun install_operations ops = Single_Assignment.assign operations ops; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
217 |
|
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
218 |
end; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
219 |
|
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
220 |
|
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
221 |
(* context-sensitive (un)checking *) |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
222 |
|
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
223 |
local |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
224 |
|
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
225 |
type key = int * bool; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
226 |
type 'a check = 'a list -> Proof.context -> ('a list * Proof.context) option; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
227 |
|
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
228 |
structure Checks = Generic_Data |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
229 |
( |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
230 |
type T = |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
231 |
((key * ((string * typ check) * stamp) list) list * |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
232 |
(key * ((string * term check) * stamp) list) list); |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
233 |
val empty = ([], []); |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
234 |
val extend = I; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
235 |
fun merge ((typ_checks1, term_checks1), (typ_checks2, term_checks2)) : T = |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
236 |
(AList.join (op =) (K (Library.merge (eq_snd (op =)))) (typ_checks1, typ_checks2), |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
237 |
AList.join (op =) (K (Library.merge (eq_snd (op =)))) (term_checks1, term_checks2)); |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
238 |
); |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
239 |
|
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
240 |
fun gen_add which (key: key) name f = |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
241 |
Checks.map (which (AList.map_default op = (key, []) (cons ((name, f), stamp ())))); |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
242 |
|
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
243 |
fun check_stage fs = perhaps_loop (perhaps_apply (map uncurry fs)); |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
244 |
|
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
245 |
fun gen_check which uncheck ctxt0 xs0 = |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
246 |
let |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
247 |
val funs = which (Checks.get (Context.Proof ctxt0)) |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
248 |
|> map_filter (fn ((i, u), fs) => if uncheck = u then SOME (i, map (snd o fst) fs) else NONE) |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
249 |
|> Library.sort (int_ord o pairself fst) |> map snd |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
250 |
|> not uncheck ? map rev; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
251 |
val check_all = perhaps_apply (map check_stage funs); |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
252 |
in #1 (perhaps check_all (xs0, ctxt0)) end; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
253 |
|
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
254 |
fun map_sort f S = |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
255 |
(case f (TFree ("", S)) of |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
256 |
TFree ("", S') => S' |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
257 |
| _ => raise TYPE ("map_sort", [TFree ("", S)], [])); |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
258 |
|
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
259 |
in |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
260 |
|
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
261 |
fun print_checks ctxt = |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
262 |
let |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
263 |
fun split_checks checks = |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
264 |
List.partition (fn ((_, un), _) => not un) checks |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
265 |
|> pairself (map (fn ((i, _), fs) => (i, map (fst o fst) fs)) |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
266 |
#> sort (int_ord o pairself fst)); |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
267 |
fun pretty_checks kind checks = |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
268 |
checks |> map (fn (i, names) => Pretty.block |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
269 |
[Pretty.str (kind ^ " (stage " ^ signed_string_of_int i ^ "):"), |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
270 |
Pretty.brk 1, Pretty.strs names]); |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
271 |
|
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
272 |
val (typs, terms) = Checks.get (Context.Proof ctxt); |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
273 |
val (typ_checks, typ_unchecks) = split_checks typs; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
274 |
val (term_checks, term_unchecks) = split_checks terms; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
275 |
in |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
276 |
pretty_checks "typ_checks" typ_checks @ |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
277 |
pretty_checks "term_checks" term_checks @ |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
278 |
pretty_checks "typ_unchecks" typ_unchecks @ |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
279 |
pretty_checks "term_unchecks" term_unchecks |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
280 |
end |> Pretty.chunks |> Pretty.writeln; |
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 |
fun add_typ_check stage = gen_add apfst (stage, false); |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
283 |
fun add_term_check stage = gen_add apsnd (stage, false); |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
284 |
fun add_typ_uncheck stage = gen_add apfst (stage, true); |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
285 |
fun add_term_uncheck stage = gen_add apsnd (stage, true); |
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 |
val check_typs = gen_check fst false; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
288 |
val check_terms = gen_check snd false; |
39288 | 289 |
fun check_props ctxt = map (Type.constraint propT) #> check_terms ctxt; |
39135
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
290 |
|
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
291 |
val check_typ = singleton o check_typs; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
292 |
val check_term = singleton o check_terms; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
293 |
val check_prop = singleton o check_props; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
294 |
val check_sort = map_sort o check_typ; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
295 |
|
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
296 |
val uncheck_typs = gen_check fst true; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
297 |
val uncheck_terms = gen_check snd true; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
298 |
val uncheck_sort = map_sort o singleton o uncheck_typs; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
299 |
|
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
300 |
end; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
301 |
|
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
302 |
|
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
303 |
(* derived operations for classrel and arity *) |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
304 |
|
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
305 |
val uncheck_classrel = map o singleton o uncheck_sort; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
306 |
|
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
307 |
fun unparse_classrel ctxt cs = Pretty.block (flat |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
308 |
(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
|
309 |
|
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
310 |
fun uncheck_arity ctxt (a, Ss, S) = |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
311 |
let |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
312 |
val T = Type (a, replicate (length Ss) dummyT); |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
313 |
val a' = |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
314 |
(case singleton (uncheck_typs ctxt) T of |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
315 |
Type (a', _) => a' |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
316 |
| T => raise TYPE ("uncheck_arity", [T], [])); |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
317 |
val Ss' = map (uncheck_sort ctxt) Ss; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
318 |
val S' = uncheck_sort ctxt S; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
319 |
in (a', Ss', S') end; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
320 |
|
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
321 |
fun unparse_arity ctxt (a, Ss, S) = |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
322 |
let |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
323 |
val prtT = unparse_typ ctxt (Type (a, [])); |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
324 |
val dom = |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
325 |
if null Ss then [] |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
326 |
else [Pretty.list "(" ")" (map (unparse_sort ctxt) Ss), Pretty.brk 1]; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
327 |
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
|
328 |
|
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
329 |
|
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
330 |
(* read = parse + check *) |
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 |
fun read_sort ctxt = parse_sort ctxt #> check_sort ctxt; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
333 |
fun read_typ ctxt = parse_typ ctxt #> singleton (check_typs ctxt); |
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 read_terms ctxt = map (parse_term ctxt) #> check_terms ctxt; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
336 |
fun read_props ctxt = map (parse_prop ctxt) #> check_props ctxt; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
337 |
|
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
338 |
val read_term = singleton o read_terms; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
339 |
val read_prop = singleton o read_props; |
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 read_sort_global = read_sort o ProofContext.init_global; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
342 |
val read_typ_global = read_typ o ProofContext.init_global; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
343 |
val read_term_global = read_term o ProofContext.init_global; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
344 |
val read_prop_global = read_prop o ProofContext.init_global; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
345 |
|
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 |
(* pretty = uncheck + unparse *) |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
348 |
|
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
349 |
fun pretty_term ctxt = singleton (uncheck_terms ctxt) #> unparse_term ctxt; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
350 |
fun pretty_typ ctxt = singleton (uncheck_typs ctxt) #> unparse_typ ctxt; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
351 |
fun pretty_sort ctxt = uncheck_sort ctxt #> unparse_sort ctxt; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
352 |
fun pretty_classrel ctxt = uncheck_classrel ctxt #> unparse_classrel ctxt; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
353 |
fun pretty_arity ctxt = uncheck_arity ctxt #> unparse_arity ctxt; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
354 |
|
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
355 |
val string_of_term = Pretty.string_of oo pretty_term; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
356 |
val string_of_typ = Pretty.string_of oo pretty_typ; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
357 |
val string_of_sort = Pretty.string_of oo pretty_sort; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
358 |
val string_of_classrel = Pretty.string_of oo pretty_classrel; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
359 |
val string_of_arity = Pretty.string_of oo pretty_arity; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
360 |
|
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
361 |
|
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
362 |
(* global pretty printing *) |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
363 |
|
39508
dfacdb01e1ec
simplified some internal flags using Config.T instead of full-blown Proof_Data;
wenzelm
parents:
39507
diff
changeset
|
364 |
val pretty_global = Config.bool (Config.declare "Syntax.pretty_global" (K (Config.Bool false))); |
dfacdb01e1ec
simplified some internal flags using Config.T instead of full-blown Proof_Data;
wenzelm
parents:
39507
diff
changeset
|
365 |
fun is_pretty_global ctxt = Config.get ctxt pretty_global; |
dfacdb01e1ec
simplified some internal flags using Config.T instead of full-blown Proof_Data;
wenzelm
parents:
39507
diff
changeset
|
366 |
val set_pretty_global = Config.put pretty_global; |
39135
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
367 |
val init_pretty_global = set_pretty_global true o ProofContext.init_global; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
368 |
|
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
369 |
val pretty_term_global = pretty_term o init_pretty_global; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
370 |
val pretty_typ_global = pretty_typ o init_pretty_global; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
371 |
val pretty_sort_global = pretty_sort o init_pretty_global; |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
372 |
|
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
373 |
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
|
374 |
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
|
375 |
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
|
376 |
|
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
377 |
|
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
378 |
(* pp operations -- deferred evaluation *) |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
379 |
|
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
380 |
fun pp ctxt = Pretty.pp |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
381 |
(fn x => pretty_term ctxt x, |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
382 |
fn x => pretty_typ ctxt x, |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
383 |
fn x => pretty_sort ctxt x, |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
384 |
fn x => pretty_classrel ctxt x, |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
385 |
fn x => pretty_arity ctxt x); |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
386 |
|
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
387 |
fun pp_global thy = Pretty.pp |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
388 |
(fn x => pretty_term_global thy x, |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
389 |
fn x => pretty_typ_global thy x, |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
390 |
fn x => pretty_sort_global thy x, |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
391 |
fn x => pretty_classrel (init_pretty_global thy) x, |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
392 |
fn x => pretty_arity (init_pretty_global thy) x); |
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
393 |
|
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
394 |
|
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
wenzelm
parents:
39126
diff
changeset
|
395 |
|
237
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
175
diff
changeset
|
396 |
(** tables of translation functions **) |
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
175
diff
changeset
|
397 |
|
5692 | 398 |
(* parse (ast) translations *) |
399 |
||
17412 | 400 |
fun lookup_tr tab c = Option.map fst (Symtab.lookup tab c); |
237
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
175
diff
changeset
|
401 |
|
23655
d2d1138e0ddc
replaced exception TableFun/GraphFun.DUPS by TableFun/GraphFun.DUP;
wenzelm
parents:
23631
diff
changeset
|
402 |
fun err_dup_trfun name c = |
d2d1138e0ddc
replaced exception TableFun/GraphFun.DUPS by TableFun/GraphFun.DUP;
wenzelm
parents:
23631
diff
changeset
|
403 |
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
|
404 |
|
37216
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
37146
diff
changeset
|
405 |
fun remove_trtab trfuns = fold (Symtab.remove Syn_Ext.eq_trfun) trfuns; |
21536
f119c730f509
extend_trtab: allow identical trfuns to be overwritten;
wenzelm
parents:
20784
diff
changeset
|
406 |
|
29004 | 407 |
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
|
408 |
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
|
409 |
|
37216
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
37146
diff
changeset
|
410 |
fun merge_trtabs name tab1 tab2 = Symtab.merge Syn_Ext.eq_trfun (tab1, tab2) |
23655
d2d1138e0ddc
replaced exception TableFun/GraphFun.DUPS by TableFun/GraphFun.DUP;
wenzelm
parents:
23631
diff
changeset
|
411 |
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
|
412 |
|
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
175
diff
changeset
|
413 |
|
5692 | 414 |
(* print (ast) translations *) |
415 |
||
18931 | 416 |
fun lookup_tr' tab c = map fst (Symtab.lookup_list tab c); |
37216
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
37146
diff
changeset
|
417 |
fun update_tr'tab trfuns = fold_rev (Symtab.update_list Syn_Ext.eq_trfun) trfuns; |
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
37146
diff
changeset
|
418 |
fun remove_tr'tab trfuns = fold (Symtab.remove_list Syn_Ext.eq_trfun) trfuns; |
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
37146
diff
changeset
|
419 |
fun merge_tr'tabs tab1 tab2 = Symtab.merge_list Syn_Ext.eq_trfun (tab1, tab2); |
5692 | 420 |
|
421 |
||
237
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
175
diff
changeset
|
422 |
|
2700 | 423 |
(** tables of token translation functions **) |
424 |
||
425 |
fun lookup_tokentr tabs modes = |
|
19482
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
wenzelm
parents:
19375
diff
changeset
|
426 |
let val trs = distinct (eq_fst (op =)) (maps (these o AList.lookup (op =) tabs) (modes @ [""])) |
17314 | 427 |
in fn c => Option.map fst (AList.lookup (op =) trs c) end; |
2700 | 428 |
|
429 |
fun merge_tokentrtabs tabs1 tabs2 = |
|
430 |
let |
|
431 |
fun eq_tr ((c1, (_, s1)), (c2, (_, s2))) = c1 = c2 andalso s1 = s2; |
|
432 |
||
4703 | 433 |
fun name (s, _) = implode (tl (Symbol.explode s)); |
2706 | 434 |
|
2700 | 435 |
fun merge mode = |
436 |
let |
|
17213 | 437 |
val trs1 = these (AList.lookup (op =) tabs1 mode); |
438 |
val trs2 = these (AList.lookup (op =) tabs2 mode); |
|
19046
bc5c6c9b114e
removed distinct, renamed gen_distinct to distinct;
wenzelm
parents:
18977
diff
changeset
|
439 |
val trs = distinct eq_tr (trs1 @ trs2); |
2700 | 440 |
in |
18964 | 441 |
(case duplicates (eq_fst (op =)) trs of |
2700 | 442 |
[] => (mode, trs) |
443 |
| dups => error ("More than one token translation function in mode " ^ |
|
2706 | 444 |
quote mode ^ " for " ^ commas_quote (map name dups))) |
2700 | 445 |
end; |
19046
bc5c6c9b114e
removed distinct, renamed gen_distinct to distinct;
wenzelm
parents:
18977
diff
changeset
|
446 |
in map merge (distinct (op =) (map fst (tabs1 @ tabs2))) end; |
2700 | 447 |
|
15755
50ac97ebe7d8
expect translations functions to be stamped already;
wenzelm
parents:
15574
diff
changeset
|
448 |
fun extend_tokentrtab tokentrs tabs = |
2700 | 449 |
let |
17213 | 450 |
fun ins_tokentr (m, c, f) = |
451 |
AList.default (op =) (m, []) |
|
452 |
#> AList.map_entry (op =) m (cons ("_" ^ c, (f, stamp ()))); |
|
15759 | 453 |
in merge_tokentrtabs tabs (fold ins_tokentr tokentrs []) end; |
2700 | 454 |
|
455 |
||
456 |
||
237
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
175
diff
changeset
|
457 |
(** tables of translation rules **) |
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
175
diff
changeset
|
458 |
|
5692 | 459 |
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
|
460 |
|
19482
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
wenzelm
parents:
19375
diff
changeset
|
461 |
fun dest_ruletab tab = maps snd (Symtab.dest tab); |
237
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
175
diff
changeset
|
462 |
|
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
175
diff
changeset
|
463 |
|
25394
db25c98f32e1
syntax operations: turned extend'' into update'' (absorb duplicates);
wenzelm
parents:
25387
diff
changeset
|
464 |
(* empty, update, merge ruletabs *) |
237
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
175
diff
changeset
|
465 |
|
25394
db25c98f32e1
syntax operations: turned extend'' into update'' (absorb duplicates);
wenzelm
parents:
25387
diff
changeset
|
466 |
val update_ruletab = fold_rev (fn r => Symtab.update_list (op =) (Ast.head_of_rule r, r)); |
18931 | 467 |
val remove_ruletab = fold (fn r => Symtab.remove_list (op =) (Ast.head_of_rule r, r)); |
468 |
fun merge_ruletabs tab1 tab2 = Symtab.merge_list (op =) (tab1, tab2); |
|
0 | 469 |
|
470 |
||
471 |
||
472 |
(** datatype syntax **) |
|
473 |
||
237
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
175
diff
changeset
|
474 |
datatype syntax = |
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
175
diff
changeset
|
475 |
Syntax of { |
37216
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
37146
diff
changeset
|
476 |
input: Syn_Ext.xprod list, |
4703 | 477 |
lexicon: Scan.lexicon, |
5692 | 478 |
gram: Parser.gram, |
18 | 479 |
consts: string list, |
2913 | 480 |
prmodes: string list, |
21772 | 481 |
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
|
482 |
parse_ruletab: ruletab, |
21772 | 483 |
parse_trtab: ((Proof.context -> term list -> term) * stamp) Symtab.table, |
484 |
print_trtab: ((Proof.context -> bool -> 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
|
485 |
print_ruletab: ruletab, |
21772 | 486 |
print_ast_trtab: ((Proof.context -> Ast.ast list -> Ast.ast) * stamp) list Symtab.table, |
26704
51ee753cc2e3
token translations: context dependent, result Pretty.T;
wenzelm
parents:
26684
diff
changeset
|
487 |
tokentrtab: (string * (string * ((Proof.context -> string -> Pretty.T) * stamp)) list) list, |
17079 | 488 |
prtabs: Printer.prtabs} * stamp; |
0 | 489 |
|
17079 | 490 |
fun eq_syntax (Syntax (_, s1), Syntax (_, s2)) = s1 = s2; |
491 |
||
492 |
fun is_keyword (Syntax ({lexicon, ...}, _)) = Scan.is_literal lexicon o Symbol.explode; |
|
14687 | 493 |
|
20784 | 494 |
type mode = string * bool; |
24970 | 495 |
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
|
496 |
val mode_input = (Print_Mode.input, true); |
15755
50ac97ebe7d8
expect translations functions to be stamped already;
wenzelm
parents:
15574
diff
changeset
|
497 |
|
18 | 498 |
|
237
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
175
diff
changeset
|
499 |
(* empty_syntax *) |
18 | 500 |
|
17079 | 501 |
val empty_syntax = Syntax |
502 |
({input = [], |
|
4703 | 503 |
lexicon = Scan.empty_lexicon, |
5692 | 504 |
gram = Parser.empty_gram, |
167 | 505 |
consts = [], |
2913 | 506 |
prmodes = [], |
5692 | 507 |
parse_ast_trtab = Symtab.empty, |
508 |
parse_ruletab = Symtab.empty, |
|
509 |
parse_trtab = Symtab.empty, |
|
510 |
print_trtab = Symtab.empty, |
|
511 |
print_ruletab = Symtab.empty, |
|
512 |
print_ast_trtab = Symtab.empty, |
|
2700 | 513 |
tokentrtab = [], |
17079 | 514 |
prtabs = Printer.empty_prtabs}, stamp ()); |
167 | 515 |
|
516 |
||
25394
db25c98f32e1
syntax operations: turned extend'' into update'' (absorb duplicates);
wenzelm
parents:
25387
diff
changeset
|
517 |
(* update_syntax *) |
167 | 518 |
|
25394
db25c98f32e1
syntax operations: turned extend'' into update'' (absorb duplicates);
wenzelm
parents:
25387
diff
changeset
|
519 |
fun update_syntax (mode, inout) syn_ext (Syntax (tabs, _)) = |
167 | 520 |
let |
15755
50ac97ebe7d8
expect translations functions to be stamped already;
wenzelm
parents:
15574
diff
changeset
|
521 |
val {input, lexicon, gram, consts = consts1, prmodes = prmodes1, |
2366 | 522 |
parse_ast_trtab, parse_ruletab, parse_trtab, print_trtab, print_ruletab, |
2700 | 523 |
print_ast_trtab, tokentrtab, prtabs} = tabs; |
37216
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
37146
diff
changeset
|
524 |
val Syn_Ext.Syn_Ext {xprods, consts = consts2, prmodes = prmodes2, |
2913 | 525 |
parse_ast_translation, parse_rules, parse_translation, print_translation, print_rules, |
2700 | 526 |
print_ast_translation, token_translation} = syn_ext; |
36208
74c5e6e3c1d3
update_syntax: add new productions only once, to allow repeated local notation, for example;
wenzelm
parents:
35668
diff
changeset
|
527 |
val new_xprods = |
74c5e6e3c1d3
update_syntax: add new productions only once, to allow repeated local notation, for example;
wenzelm
parents:
35668
diff
changeset
|
528 |
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
|
529 |
fun if_inout xs = if inout then xs else []; |
167 | 530 |
in |
17079 | 531 |
Syntax |
36208
74c5e6e3c1d3
update_syntax: add new productions only once, to allow repeated local notation, for example;
wenzelm
parents:
35668
diff
changeset
|
532 |
({input = new_xprods @ input, |
37216
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
37146
diff
changeset
|
533 |
lexicon = fold Scan.extend_lexicon (Syn_Ext.delims_of new_xprods) lexicon, |
37684 | 534 |
gram = Parser.extend_gram new_xprods gram, |
35429
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35412
diff
changeset
|
535 |
consts = Library.merge (op =) (consts1, filter_out Lexicon.is_marked consts2), |
18921 | 536 |
prmodes = insert (op =) mode (Library.merge (op =) (prmodes1, prmodes2)), |
167 | 537 |
parse_ast_trtab = |
25394
db25c98f32e1
syntax operations: turned extend'' into update'' (absorb duplicates);
wenzelm
parents:
25387
diff
changeset
|
538 |
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
|
539 |
parse_ruletab = update_ruletab (if_inout parse_rules) parse_ruletab, |
db25c98f32e1
syntax operations: turned extend'' into update'' (absorb duplicates);
wenzelm
parents:
25387
diff
changeset
|
540 |
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
|
541 |
print_trtab = update_tr'tab print_translation print_trtab, |
db25c98f32e1
syntax operations: turned extend'' into update'' (absorb duplicates);
wenzelm
parents:
25387
diff
changeset
|
542 |
print_ruletab = update_ruletab print_rules print_ruletab, |
db25c98f32e1
syntax operations: turned extend'' into update'' (absorb duplicates);
wenzelm
parents:
25387
diff
changeset
|
543 |
print_ast_trtab = update_tr'tab print_ast_translation print_ast_trtab, |
15755
50ac97ebe7d8
expect translations functions to be stamped already;
wenzelm
parents:
15574
diff
changeset
|
544 |
tokentrtab = extend_tokentrtab token_translation tokentrtab, |
25394
db25c98f32e1
syntax operations: turned extend'' into update'' (absorb duplicates);
wenzelm
parents:
25387
diff
changeset
|
545 |
prtabs = Printer.update_prtabs mode xprods prtabs}, stamp ()) |
18 | 546 |
end; |
547 |
||
15755
50ac97ebe7d8
expect translations functions to be stamped already;
wenzelm
parents:
15574
diff
changeset
|
548 |
|
50ac97ebe7d8
expect translations functions to be stamped already;
wenzelm
parents:
15574
diff
changeset
|
549 |
(* remove_syntax *) |
50ac97ebe7d8
expect translations functions to be stamped already;
wenzelm
parents:
15574
diff
changeset
|
550 |
|
17079 | 551 |
fun remove_syntax (mode, inout) syn_ext (Syntax (tabs, _)) = |
15755
50ac97ebe7d8
expect translations functions to be stamped already;
wenzelm
parents:
15574
diff
changeset
|
552 |
let |
37216
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
37146
diff
changeset
|
553 |
val Syn_Ext.Syn_Ext {xprods, consts = _, prmodes = _, |
15755
50ac97ebe7d8
expect translations functions to be stamped already;
wenzelm
parents:
15574
diff
changeset
|
554 |
parse_ast_translation, parse_rules, parse_translation, print_translation, print_rules, |
50ac97ebe7d8
expect translations functions to be stamped already;
wenzelm
parents:
15574
diff
changeset
|
555 |
print_ast_translation, token_translation = _} = syn_ext; |
50ac97ebe7d8
expect translations functions to be stamped already;
wenzelm
parents:
15574
diff
changeset
|
556 |
val {input, lexicon, gram, consts, prmodes, |
50ac97ebe7d8
expect translations functions to be stamped already;
wenzelm
parents:
15574
diff
changeset
|
557 |
parse_ast_trtab, parse_ruletab, parse_trtab, print_trtab, print_ruletab, |
50ac97ebe7d8
expect translations functions to be stamped already;
wenzelm
parents:
15574
diff
changeset
|
558 |
print_ast_trtab, tokentrtab, prtabs} = tabs; |
19300 | 559 |
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
|
560 |
val changed = length input <> length input'; |
19546
00d5c7c7ce07
extend/remove_syntax: observe inout flag for translations, too;
wenzelm
parents:
19482
diff
changeset
|
561 |
fun if_inout xs = if inout then xs else []; |
15755
50ac97ebe7d8
expect translations functions to be stamped already;
wenzelm
parents:
15574
diff
changeset
|
562 |
in |
17079 | 563 |
Syntax |
564 |
({input = input', |
|
37216
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
37146
diff
changeset
|
565 |
lexicon = if changed then Scan.make_lexicon (Syn_Ext.delims_of input') else lexicon, |
25387 | 566 |
gram = if changed then Parser.make_gram input' else gram, |
15755
50ac97ebe7d8
expect translations functions to be stamped already;
wenzelm
parents:
15574
diff
changeset
|
567 |
consts = consts, |
50ac97ebe7d8
expect translations functions to be stamped already;
wenzelm
parents:
15574
diff
changeset
|
568 |
prmodes = prmodes, |
19546
00d5c7c7ce07
extend/remove_syntax: observe inout flag for translations, too;
wenzelm
parents:
19482
diff
changeset
|
569 |
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
|
570 |
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
|
571 |
parse_trtab = remove_trtab (if_inout parse_translation) parse_trtab, |
15755
50ac97ebe7d8
expect translations functions to be stamped already;
wenzelm
parents:
15574
diff
changeset
|
572 |
print_trtab = remove_tr'tab print_translation print_trtab, |
50ac97ebe7d8
expect translations functions to be stamped already;
wenzelm
parents:
15574
diff
changeset
|
573 |
print_ruletab = remove_ruletab print_rules print_ruletab, |
50ac97ebe7d8
expect translations functions to be stamped already;
wenzelm
parents:
15574
diff
changeset
|
574 |
print_ast_trtab = remove_tr'tab print_ast_translation print_ast_trtab, |
50ac97ebe7d8
expect translations functions to be stamped already;
wenzelm
parents:
15574
diff
changeset
|
575 |
tokentrtab = tokentrtab, |
17079 | 576 |
prtabs = Printer.remove_prtabs mode xprods prtabs}, stamp ()) |
15755
50ac97ebe7d8
expect translations functions to be stamped already;
wenzelm
parents:
15574
diff
changeset
|
577 |
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
|
578 |
|
18 | 579 |
|
237
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
175
diff
changeset
|
580 |
(* merge_syntaxes *) |
0 | 581 |
|
17079 | 582 |
fun merge_syntaxes (Syntax (tabs1, _)) (Syntax (tabs2, _)) = |
0 | 583 |
let |
15755
50ac97ebe7d8
expect translations functions to be stamped already;
wenzelm
parents:
15574
diff
changeset
|
584 |
val {input = input1, lexicon = lexicon1, gram = gram1, consts = consts1, |
50ac97ebe7d8
expect translations functions to be stamped already;
wenzelm
parents:
15574
diff
changeset
|
585 |
prmodes = prmodes1, parse_ast_trtab = parse_ast_trtab1, |
50ac97ebe7d8
expect translations functions to be stamped already;
wenzelm
parents:
15574
diff
changeset
|
586 |
parse_ruletab = parse_ruletab1, parse_trtab = parse_trtab1, |
50ac97ebe7d8
expect translations functions to be stamped already;
wenzelm
parents:
15574
diff
changeset
|
587 |
print_trtab = print_trtab1, print_ruletab = print_ruletab1, |
50ac97ebe7d8
expect translations functions to be stamped already;
wenzelm
parents:
15574
diff
changeset
|
588 |
print_ast_trtab = print_ast_trtab1, tokentrtab = tokentrtab1, prtabs = prtabs1} = tabs1; |
237
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
175
diff
changeset
|
589 |
|
15755
50ac97ebe7d8
expect translations functions to be stamped already;
wenzelm
parents:
15574
diff
changeset
|
590 |
val {input = input2, lexicon = lexicon2, gram = gram2, consts = consts2, |
50ac97ebe7d8
expect translations functions to be stamped already;
wenzelm
parents:
15574
diff
changeset
|
591 |
prmodes = prmodes2, parse_ast_trtab = parse_ast_trtab2, |
50ac97ebe7d8
expect translations functions to be stamped already;
wenzelm
parents:
15574
diff
changeset
|
592 |
parse_ruletab = parse_ruletab2, parse_trtab = parse_trtab2, |
50ac97ebe7d8
expect translations functions to be stamped already;
wenzelm
parents:
15574
diff
changeset
|
593 |
print_trtab = print_trtab2, print_ruletab = print_ruletab2, |
50ac97ebe7d8
expect translations functions to be stamped already;
wenzelm
parents:
15574
diff
changeset
|
594 |
print_ast_trtab = print_ast_trtab2, tokentrtab = tokentrtab2, prtabs = prtabs2} = tabs2; |
0 | 595 |
in |
17079 | 596 |
Syntax |
18921 | 597 |
({input = Library.merge (op =) (input1, input2), |
27768 | 598 |
lexicon = Scan.merge_lexicons (lexicon1, lexicon2), |
37684 | 599 |
gram = Parser.merge_gram (gram1, gram2), |
18428 | 600 |
consts = sort_distinct string_ord (consts1 @ consts2), |
18921 | 601 |
prmodes = Library.merge (op =) (prmodes1, prmodes2), |
237
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
175
diff
changeset
|
602 |
parse_ast_trtab = |
15755
50ac97ebe7d8
expect translations functions to be stamped already;
wenzelm
parents:
15574
diff
changeset
|
603 |
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
|
604 |
parse_ruletab = merge_ruletabs parse_ruletab1 parse_ruletab2, |
15755
50ac97ebe7d8
expect translations functions to be stamped already;
wenzelm
parents:
15574
diff
changeset
|
605 |
parse_trtab = merge_trtabs "parse translation" parse_trtab1 parse_trtab2, |
5692 | 606 |
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
|
607 |
print_ruletab = merge_ruletabs print_ruletab1 print_ruletab2, |
5692 | 608 |
print_ast_trtab = merge_tr'tabs print_ast_trtab1 print_ast_trtab2, |
2700 | 609 |
tokentrtab = merge_tokentrtabs tokentrtab1 tokentrtab2, |
17079 | 610 |
prtabs = Printer.merge_prtabs prtabs1 prtabs2}, stamp ()) |
0 | 611 |
end; |
612 |
||
613 |
||
18720 | 614 |
(* basic syntax *) |
260 | 615 |
|
35668 | 616 |
val basic_syntax = |
17168 | 617 |
empty_syntax |
37216
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
37146
diff
changeset
|
618 |
|> update_syntax mode_default Type_Ext.type_ext |
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
37146
diff
changeset
|
619 |
|> update_syntax mode_default Syn_Ext.pure_ext; |
260 | 620 |
|
18720 | 621 |
val basic_nonterms = |
37216
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
37146
diff
changeset
|
622 |
(Lexicon.terminals @ [Syn_Ext.logic, "type", "types", "sort", "classes", |
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
37146
diff
changeset
|
623 |
Syn_Ext.args, Syn_Ext.cargs, "pttrn", "pttrns", "idt", "idts", "aprop", |
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
37146
diff
changeset
|
624 |
"asms", Syn_Ext.any, Syn_Ext.sprop, "num_const", "float_const", |
28904 | 625 |
"index", "struct"]); |
18720 | 626 |
|
0 | 627 |
|
4887 | 628 |
|
15759 | 629 |
(** print syntax **) |
630 |
||
631 |
local |
|
0 | 632 |
|
260 | 633 |
fun pretty_strs_qs name strs = |
28840 | 634 |
Pretty.strs (name :: map quote (sort_strings strs)); |
0 | 635 |
|
17079 | 636 |
fun pretty_gram (Syntax (tabs, _)) = |
0 | 637 |
let |
32784 | 638 |
val {lexicon, prmodes, gram, ...} = tabs; |
28375 | 639 |
val prmodes' = sort_strings (filter_out (fn s => s = "") prmodes); |
0 | 640 |
in |
8720 | 641 |
[pretty_strs_qs "lexicon:" (Scan.dest_lexicon lexicon), |
642 |
Pretty.big_list "prods:" (Parser.pretty_gram gram), |
|
643 |
pretty_strs_qs "print modes:" prmodes'] |
|
0 | 644 |
end; |
645 |
||
17079 | 646 |
fun pretty_trans (Syntax (tabs, _)) = |
0 | 647 |
let |
260 | 648 |
fun pretty_trtab name tab = |
5692 | 649 |
pretty_strs_qs name (Symtab.keys tab); |
0 | 650 |
|
260 | 651 |
fun pretty_ruletab name tab = |
5692 | 652 |
Pretty.big_list name (map Ast.pretty_rule (dest_ruletab tab)); |
0 | 653 |
|
28840 | 654 |
fun pretty_tokentr (mode, trs) = Pretty.strs (quote mode ^ ":" :: map fst trs); |
4703 | 655 |
|
237
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
175
diff
changeset
|
656 |
val {consts, parse_ast_trtab, parse_ruletab, parse_trtab, print_trtab, |
4703 | 657 |
print_ruletab, print_ast_trtab, tokentrtab, ...} = tabs; |
0 | 658 |
in |
8720 | 659 |
[pretty_strs_qs "consts:" consts, |
660 |
pretty_trtab "parse_ast_translation:" parse_ast_trtab, |
|
661 |
pretty_ruletab "parse_rules:" parse_ruletab, |
|
662 |
pretty_trtab "parse_translation:" parse_trtab, |
|
663 |
pretty_trtab "print_translation:" print_trtab, |
|
664 |
pretty_ruletab "print_rules:" print_ruletab, |
|
665 |
pretty_trtab "print_ast_translation:" print_ast_trtab, |
|
666 |
Pretty.big_list "token_translation:" (map pretty_tokentr tokentrtab)] |
|
0 | 667 |
end; |
668 |
||
15759 | 669 |
in |
0 | 670 |
|
15759 | 671 |
fun print_gram syn = Pretty.writeln (Pretty.chunks (pretty_gram syn)); |
672 |
fun print_trans syn = Pretty.writeln (Pretty.chunks (pretty_trans syn)); |
|
673 |
fun print_syntax syn = Pretty.writeln (Pretty.chunks (pretty_gram syn @ pretty_trans syn)); |
|
0 | 674 |
|
15759 | 675 |
end; |
0 | 676 |
|
677 |
||
26951
030e4a818b39
moved global pretty/string_of functions from Sign to Syntax;
wenzelm
parents:
26704
diff
changeset
|
678 |
(* reconstructing infixes -- educated guessing *) |
030e4a818b39
moved global pretty/string_of functions from Sign to Syntax;
wenzelm
parents:
26704
diff
changeset
|
679 |
|
030e4a818b39
moved global pretty/string_of functions from Sign to Syntax;
wenzelm
parents:
26704
diff
changeset
|
680 |
fun guess_infix (Syntax ({gram, ...}, _)) c = |
030e4a818b39
moved global pretty/string_of functions from Sign to Syntax;
wenzelm
parents:
26704
diff
changeset
|
681 |
(case Parser.guess_infix_lr gram c of |
030e4a818b39
moved global pretty/string_of functions from Sign to Syntax;
wenzelm
parents:
26704
diff
changeset
|
682 |
SOME (s, l, r, j) => SOME |
35130 | 683 |
(if l then Mixfix.Infixl (s, j) |
684 |
else if r then Mixfix.Infixr (s, j) |
|
685 |
else Mixfix.Infix (s, j)) |
|
26951
030e4a818b39
moved global pretty/string_of functions from Sign to Syntax;
wenzelm
parents:
26704
diff
changeset
|
686 |
| NONE => NONE); |
030e4a818b39
moved global pretty/string_of functions from Sign to Syntax;
wenzelm
parents:
26704
diff
changeset
|
687 |
|
030e4a818b39
moved global pretty/string_of functions from Sign to Syntax;
wenzelm
parents:
26704
diff
changeset
|
688 |
|
0 | 689 |
|
237
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
175
diff
changeset
|
690 |
(** read **) |
18 | 691 |
|
260 | 692 |
(* read_ast *) |
693 |
||
39126
ee117c5b3b75
configuration options Syntax.ambiguity_enabled (inverse of former Syntax.ambiguity_is_error), Syntax.ambiguity_level (with Isar attribute "syntax_ambiguity_level"), Syntax.ambiguity_limit;
wenzelm
parents:
38831
diff
changeset
|
694 |
val ambiguity_enabled = |
ee117c5b3b75
configuration options Syntax.ambiguity_enabled (inverse of former Syntax.ambiguity_is_error), Syntax.ambiguity_level (with Isar attribute "syntax_ambiguity_level"), Syntax.ambiguity_limit;
wenzelm
parents:
38831
diff
changeset
|
695 |
Config.bool (Config.declare "syntax_ambiguity_enabled" (fn _ => Config.Bool true)); |
ee117c5b3b75
configuration options Syntax.ambiguity_enabled (inverse of former Syntax.ambiguity_is_error), Syntax.ambiguity_level (with Isar attribute "syntax_ambiguity_level"), Syntax.ambiguity_limit;
wenzelm
parents:
38831
diff
changeset
|
696 |
|
39163
4d701c0388c3
more explicit indication of Config.raw options, which are only needed for bootstrapping Pure;
wenzelm
parents:
39137
diff
changeset
|
697 |
val ambiguity_level_raw = Config.declare "syntax_ambiguity_level" (fn _ => Config.Int 1); |
4d701c0388c3
more explicit indication of Config.raw options, which are only needed for bootstrapping Pure;
wenzelm
parents:
39137
diff
changeset
|
698 |
val ambiguity_level = Config.int ambiguity_level_raw; |
39126
ee117c5b3b75
configuration options Syntax.ambiguity_enabled (inverse of former Syntax.ambiguity_is_error), Syntax.ambiguity_level (with Isar attribute "syntax_ambiguity_level"), Syntax.ambiguity_limit;
wenzelm
parents:
38831
diff
changeset
|
699 |
|
ee117c5b3b75
configuration options Syntax.ambiguity_enabled (inverse of former Syntax.ambiguity_is_error), Syntax.ambiguity_level (with Isar attribute "syntax_ambiguity_level"), Syntax.ambiguity_limit;
wenzelm
parents:
38831
diff
changeset
|
700 |
val ambiguity_limit = |
ee117c5b3b75
configuration options Syntax.ambiguity_enabled (inverse of former Syntax.ambiguity_is_error), Syntax.ambiguity_level (with Isar attribute "syntax_ambiguity_level"), Syntax.ambiguity_limit;
wenzelm
parents:
38831
diff
changeset
|
701 |
Config.int (Config.declare "syntax_ambiguity_limit" (fn _ => Config.Int 10)); |
882
b118d1ea0dfd
moved ambiguity_level from sign.ML to Syntax/syntax.ML
clasohm
parents:
865
diff
changeset
|
702 |
|
27786
4525c6f5d753
added read_token -- with optional YXML encoding of position;
wenzelm
parents:
27768
diff
changeset
|
703 |
fun ambiguity_msg pos = "Parse error: ambiguous syntax" ^ Position.str_of pos; |
4525c6f5d753
added read_token -- with optional YXML encoding of position;
wenzelm
parents:
27768
diff
changeset
|
704 |
|
4525c6f5d753
added read_token -- with optional YXML encoding of position;
wenzelm
parents:
27768
diff
changeset
|
705 |
fun read_asts ctxt is_logtype (Syntax (tabs, _)) xids root (syms, pos) = |
260 | 706 |
let |
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
|
707 |
val {lexicon, gram, parse_ast_trtab, ...} = tabs; |
27889
c9e8a5bda32b
read_asts: Lexicon.report_token, filter Lexicon.is_proper;
wenzelm
parents:
27822
diff
changeset
|
708 |
val toks = Lexicon.tokenize lexicon xids syms; |
38237
8b0383334031
prefer Context_Position.report where a proper context is available -- notably for "inner" entities;
wenzelm
parents:
38229
diff
changeset
|
709 |
val _ = List.app (Lexicon.report_token ctxt) toks; |
27889
c9e8a5bda32b
read_asts: Lexicon.report_token, filter Lexicon.is_proper;
wenzelm
parents:
27822
diff
changeset
|
710 |
|
37216
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
37146
diff
changeset
|
711 |
val root' = if root <> "prop" andalso is_logtype root then Syn_Ext.logic else root; |
39168
e3ac771235f7
report token range after inner parse error -- often provides important clues about misunderstanding concerning lexical phase;
wenzelm
parents:
39163
diff
changeset
|
712 |
val pts = Parser.parse ctxt gram root' (filter Lexicon.is_proper toks) |
39510
d9f5f01faa1b
Syntax.read_asts error: report token ranges within message -- no side-effect here;
wenzelm
parents:
39508
diff
changeset
|
713 |
handle ERROR msg => |
d9f5f01faa1b
Syntax.read_asts error: report token ranges within message -- no side-effect here;
wenzelm
parents:
39508
diff
changeset
|
714 |
error (msg ^ |
d9f5f01faa1b
Syntax.read_asts error: report token ranges within message -- no side-effect here;
wenzelm
parents:
39508
diff
changeset
|
715 |
implode (map (Markup.markup Markup.report o Lexicon.reported_token_range ctxt) toks)); |
25993
d542486e39e7
added ambiguity_limit (restricts parse trees / terms printed in messages);
wenzelm
parents:
25476
diff
changeset
|
716 |
val len = length pts; |
330
2fda15dd1e0f
changed the way a grammar is generated to allow the new parser to work;
clasohm
parents:
260
diff
changeset
|
717 |
|
39126
ee117c5b3b75
configuration options Syntax.ambiguity_enabled (inverse of former Syntax.ambiguity_is_error), Syntax.ambiguity_level (with Isar attribute "syntax_ambiguity_level"), Syntax.ambiguity_limit;
wenzelm
parents:
38831
diff
changeset
|
718 |
val limit = Config.get ctxt ambiguity_limit; |
12292 | 719 |
fun show_pt pt = |
37216
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
37146
diff
changeset
|
720 |
Pretty.string_of (Ast.pretty_ast (hd (Syn_Trans.pts_to_asts ctxt (K NONE) [pt]))); |
260 | 721 |
in |
39126
ee117c5b3b75
configuration options Syntax.ambiguity_enabled (inverse of former Syntax.ambiguity_is_error), Syntax.ambiguity_level (with Isar attribute "syntax_ambiguity_level"), Syntax.ambiguity_limit;
wenzelm
parents:
38831
diff
changeset
|
722 |
if len <= Config.get ctxt ambiguity_level then () |
ee117c5b3b75
configuration options Syntax.ambiguity_enabled (inverse of former Syntax.ambiguity_is_error), Syntax.ambiguity_level (with Isar attribute "syntax_ambiguity_level"), Syntax.ambiguity_limit;
wenzelm
parents:
38831
diff
changeset
|
723 |
else if not (Config.get ctxt ambiguity_enabled) then error (ambiguity_msg pos) |
25993
d542486e39e7
added ambiguity_limit (restricts parse trees / terms printed in messages);
wenzelm
parents:
25476
diff
changeset
|
724 |
else |
38831
4933a3dfd745
more careful treatment of context visibility flag wrt. spurious warnings;
wenzelm
parents:
38238
diff
changeset
|
725 |
(Context_Position.if_visible ctxt warning (cat_lines |
33919 | 726 |
(("Ambiguous input" ^ Position.str_of pos ^ |
27786
4525c6f5d753
added read_token -- with optional YXML encoding of position;
wenzelm
parents:
27768
diff
changeset
|
727 |
"\nproduces " ^ string_of_int len ^ " parse trees" ^ |
25993
d542486e39e7
added ambiguity_limit (restricts parse trees / terms printed in messages);
wenzelm
parents:
25476
diff
changeset
|
728 |
(if len <= limit then "" else " (" ^ string_of_int limit ^ " displayed)") ^ ":") :: |
33957 | 729 |
map show_pt (take limit pts)))); |
37216
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
37146
diff
changeset
|
730 |
Syn_Trans.pts_to_asts ctxt (lookup_tr parse_ast_trtab) pts |
260 | 731 |
end; |
732 |
||
733 |
||
237
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
175
diff
changeset
|
734 |
(* read *) |
0 | 735 |
|
27786
4525c6f5d753
added read_token -- with optional YXML encoding of position;
wenzelm
parents:
27768
diff
changeset
|
736 |
fun read ctxt is_logtype (syn as Syntax (tabs, _)) ty inp = |
0 | 737 |
let |
237
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
175
diff
changeset
|
738 |
val {parse_ruletab, parse_trtab, ...} = tabs; |
37216
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
37146
diff
changeset
|
739 |
val asts = read_asts ctxt is_logtype syn false (Syn_Ext.typ_to_nonterm ty) inp; |
0 | 740 |
in |
37216
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
37146
diff
changeset
|
741 |
Syn_Trans.asts_to_terms ctxt (lookup_tr parse_trtab) |
18931 | 742 |
(map (Ast.normalize_ast (Symtab.lookup_list parse_ruletab)) asts) |
0 | 743 |
end; |
744 |
||
745 |
||
22703 | 746 |
(* read terms *) |
747 |
||
24372
743575ccfec8
standard_parse_term: added pp/check argument, include disambig here (from sign.ML);
wenzelm
parents:
24341
diff
changeset
|
748 |
(*brute-force disambiguation via type-inference*) |
39136
b0b3b6318e3b
Syntax.standard_parse_term: eliminated redundant Pretty.pp;
wenzelm
parents:
39135
diff
changeset
|
749 |
fun disambig _ _ [t] = t |
b0b3b6318e3b
Syntax.standard_parse_term: eliminated redundant Pretty.pp;
wenzelm
parents:
39135
diff
changeset
|
750 |
| disambig ctxt check ts = |
24372
743575ccfec8
standard_parse_term: added pp/check argument, include disambig here (from sign.ML);
wenzelm
parents:
24341
diff
changeset
|
751 |
let |
39126
ee117c5b3b75
configuration options Syntax.ambiguity_enabled (inverse of former Syntax.ambiguity_is_error), Syntax.ambiguity_level (with Isar attribute "syntax_ambiguity_level"), Syntax.ambiguity_limit;
wenzelm
parents:
38831
diff
changeset
|
752 |
val level = Config.get ctxt ambiguity_level; |
ee117c5b3b75
configuration options Syntax.ambiguity_enabled (inverse of former Syntax.ambiguity_is_error), Syntax.ambiguity_level (with Isar attribute "syntax_ambiguity_level"), Syntax.ambiguity_limit;
wenzelm
parents:
38831
diff
changeset
|
753 |
val limit = Config.get ctxt ambiguity_limit; |
24372
743575ccfec8
standard_parse_term: added pp/check argument, include disambig here (from sign.ML);
wenzelm
parents:
24341
diff
changeset
|
754 |
|
743575ccfec8
standard_parse_term: added pp/check argument, include disambig here (from sign.ML);
wenzelm
parents:
24341
diff
changeset
|
755 |
val ambiguity = length ts; |
743575ccfec8
standard_parse_term: added pp/check argument, include disambig here (from sign.ML);
wenzelm
parents:
24341
diff
changeset
|
756 |
fun ambig_msg () = |
25993
d542486e39e7
added ambiguity_limit (restricts parse trees / terms printed in messages);
wenzelm
parents:
25476
diff
changeset
|
757 |
if ambiguity > 1 andalso ambiguity <= level then |
24372
743575ccfec8
standard_parse_term: added pp/check argument, include disambig here (from sign.ML);
wenzelm
parents:
24341
diff
changeset
|
758 |
"Got more than one parse tree.\n\ |
39126
ee117c5b3b75
configuration options Syntax.ambiguity_enabled (inverse of former Syntax.ambiguity_is_error), Syntax.ambiguity_level (with Isar attribute "syntax_ambiguity_level"), Syntax.ambiguity_limit;
wenzelm
parents:
38831
diff
changeset
|
759 |
\Retry with smaller syntax_ambiguity_level for more information." |
24372
743575ccfec8
standard_parse_term: added pp/check argument, include disambig here (from sign.ML);
wenzelm
parents:
24341
diff
changeset
|
760 |
else ""; |
25993
d542486e39e7
added ambiguity_limit (restricts parse trees / terms printed in messages);
wenzelm
parents:
25476
diff
changeset
|
761 |
|
35394 | 762 |
val errs = Par_List.map check ts; |
25993
d542486e39e7
added ambiguity_limit (restricts parse trees / terms printed in messages);
wenzelm
parents:
25476
diff
changeset
|
763 |
val results = map_filter (fn (t, NONE) => SOME t | _ => NONE) (ts ~~ errs); |
d542486e39e7
added ambiguity_limit (restricts parse trees / terms printed in messages);
wenzelm
parents:
25476
diff
changeset
|
764 |
val len = length results; |
39136
b0b3b6318e3b
Syntax.standard_parse_term: eliminated redundant Pretty.pp;
wenzelm
parents:
39135
diff
changeset
|
765 |
|
39137
ccb53edd59f0
turned show_brackets into proper configuration option;
wenzelm
parents:
39136
diff
changeset
|
766 |
val show_term = string_of_term (Config.put Printer.show_brackets true ctxt); |
24372
743575ccfec8
standard_parse_term: added pp/check argument, include disambig here (from sign.ML);
wenzelm
parents:
24341
diff
changeset
|
767 |
in |
25476
03da46cfab9e
standard_parse_term: check ambiguous results without changing the result yet;
wenzelm
parents:
25394
diff
changeset
|
768 |
if null results then cat_error (ambig_msg ()) (cat_lines (map_filter I errs)) |
25993
d542486e39e7
added ambiguity_limit (restricts parse trees / terms printed in messages);
wenzelm
parents:
25476
diff
changeset
|
769 |
else if len = 1 then |
d542486e39e7
added ambiguity_limit (restricts parse trees / terms printed in messages);
wenzelm
parents:
25476
diff
changeset
|
770 |
(if ambiguity > level then |
38831
4933a3dfd745
more careful treatment of context visibility flag wrt. spurious warnings;
wenzelm
parents:
38238
diff
changeset
|
771 |
Context_Position.if_visible ctxt warning |
4933a3dfd745
more careful treatment of context visibility flag wrt. spurious warnings;
wenzelm
parents:
38238
diff
changeset
|
772 |
"Fortunately, only one parse tree is type correct.\n\ |
24372
743575ccfec8
standard_parse_term: added pp/check argument, include disambig here (from sign.ML);
wenzelm
parents:
24341
diff
changeset
|
773 |
\You may still want to disambiguate your grammar or your input." |
743575ccfec8
standard_parse_term: added pp/check argument, include disambig here (from sign.ML);
wenzelm
parents:
24341
diff
changeset
|
774 |
else (); hd results) |
25993
d542486e39e7
added ambiguity_limit (restricts parse trees / terms printed in messages);
wenzelm
parents:
25476
diff
changeset
|
775 |
else cat_error (ambig_msg ()) (cat_lines |
d542486e39e7
added ambiguity_limit (restricts parse trees / terms printed in messages);
wenzelm
parents:
25476
diff
changeset
|
776 |
(("Ambiguous input, " ^ string_of_int len ^ " terms are type correct" ^ |
d542486e39e7
added ambiguity_limit (restricts parse trees / terms printed in messages);
wenzelm
parents:
25476
diff
changeset
|
777 |
(if len <= limit then "" else " (" ^ string_of_int limit ^ " displayed)") ^ ":") :: |
39136
b0b3b6318e3b
Syntax.standard_parse_term: eliminated redundant Pretty.pp;
wenzelm
parents:
39135
diff
changeset
|
778 |
map show_term (take limit results))) |
24372
743575ccfec8
standard_parse_term: added pp/check argument, include disambig here (from sign.ML);
wenzelm
parents:
24341
diff
changeset
|
779 |
end; |
743575ccfec8
standard_parse_term: added pp/check argument, include disambig here (from sign.ML);
wenzelm
parents:
24341
diff
changeset
|
780 |
|
39136
b0b3b6318e3b
Syntax.standard_parse_term: eliminated redundant Pretty.pp;
wenzelm
parents:
39135
diff
changeset
|
781 |
fun standard_parse_term check get_sort map_const map_free ctxt is_logtype syn ty (syms, pos) = |
27786
4525c6f5d753
added read_token -- with optional YXML encoding of position;
wenzelm
parents:
27768
diff
changeset
|
782 |
read ctxt is_logtype syn ty (syms, pos) |
37216
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
37146
diff
changeset
|
783 |
|> map (Type_Ext.decode_term get_sort map_const map_free) |
39136
b0b3b6318e3b
Syntax.standard_parse_term: eliminated redundant Pretty.pp;
wenzelm
parents:
39135
diff
changeset
|
784 |
|> disambig ctxt check; |
22703 | 785 |
|
786 |
||
237
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents:
175
diff
changeset
|
787 |
(* read types *) |
0 | 788 |
|
35429
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35412
diff
changeset
|
789 |
fun standard_parse_typ ctxt syn get_sort (syms, pos) = |
37216
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
37146
diff
changeset
|
790 |
(case read ctxt (K false) syn Syn_Ext.typeT (syms, pos) of |
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
37146
diff
changeset
|
791 |
[t] => Type_Ext.typ_of_term (get_sort (Type_Ext.term_sorts t)) t |
27786
4525c6f5d753
added read_token -- with optional YXML encoding of position;
wenzelm
parents:
27768
diff
changeset
|
792 |
| _ => error (ambiguity_msg pos)); |
144 | 793 |
|
794 |
||
8894 | 795 |
(* read sorts *) |
796 |
||
35429
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35412
diff
changeset
|
797 |
fun standard_parse_sort ctxt syn (syms, pos) = |
37216
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
37146
diff
changeset
|
798 |
(case read ctxt (K false) syn Type_Ext.sortT (syms, pos) of |
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
37146
diff
changeset
|
799 |
[t] => Type_Ext.sort_of_term t |
27786
4525c6f5d753
added read_token -- with optional YXML encoding of position;
wenzelm
parents:
27768
diff
changeset
|
800 |
| _ => error (ambiguity_msg pos)); |
8894 | 801 |
|
802 |
||
18 | 803 |
|
1158 | 804 |
(** prepare translation rules **) |
805 |
||
806 |
datatype 'a trrule = |
|
3526 | 807 |
ParseRule of 'a * 'a | |
808 |
PrintRule of 'a * 'a | |
|
809 |
ParsePrintRule of 'a * 'a; |
|
888
3a1de9454d13
improved read_xrules: patterns no longer read twice;
wenzelm
parents:
882
diff
changeset
|
810 |
|
4618 | 811 |
fun map_trrule f (ParseRule (x, y)) = ParseRule (f x, f y) |
812 |
| map_trrule f (PrintRule (x, y)) = PrintRule (f x, f y) |
|
813 |
| map_trrule f (ParsePrintRule (x, y)) = ParsePrintRule (f x, f y); |
|
1158 | 814 |
|
15531 | 815 |
fun parse_rule (ParseRule pats) = SOME pats |
816 |
| parse_rule (PrintRule _) = NONE |
|
817 |
| parse_rule (ParsePrintRule pats) = SOME pats; |
|
1158 | 818 |
|
15531 | 819 |
fun print_rule (ParseRule _) = NONE |
820 |
| print_rule (PrintRule pats) = SOME (swap pats) |
|
821 |
| print_rule (ParsePrintRule pats) = SOME (swap pats); |
|
1158 | 822 |
|
823 |
||
35111 | 824 |
fun is_const (Syntax ({consts, ...}, _)) c = member (op =) consts c; |
825 |
||
19262 | 826 |
local |
827 |
||
1158 | 828 |
fun check_rule (rule as (lhs, rhs)) = |
5692 | 829 |
(case Ast.rule_error rule of |
15531 | 830 |
SOME msg => |
1158 | 831 |
error ("Error in syntax translation rule: " ^ msg ^ "\n" ^ |
5692 | 832 |
Ast.str_of_ast lhs ^ " -> " ^ Ast.str_of_ast rhs) |
15531 | 833 |
| NONE => rule); |
888
3a1de9454d13
improved read_xrules: patterns no longer read twice;
wenzelm
parents:
882
diff
changeset
|
834 |
|
21772 | 835 |
fun read_pattern ctxt is_logtype syn (root, str) = |
18 | 836 |
let |
5692 | 837 |
fun constify (ast as Ast.Constant _) = ast |
838 |
| constify (ast as Ast.Variable x) = |
|
35111 | 839 |
if is_const syn x orelse Long_Name.is_qualified x then Ast.Constant x |
3830 | 840 |
else ast |
5692 | 841 |
| constify (Ast.Appl asts) = Ast.Appl (map constify asts); |
27786
4525c6f5d753
added read_token -- with optional YXML encoding of position;
wenzelm
parents:
27768
diff
changeset
|
842 |
|
4525c6f5d753
added read_token -- with optional YXML encoding of position;
wenzelm
parents:
27768
diff
changeset
|
843 |
val (syms, pos) = read_token str; |
18 | 844 |
in |
27786
4525c6f5d753
added read_token -- with optional YXML encoding of position;
wenzelm
parents:
27768
diff
changeset
|
845 |
(case read_asts ctxt is_logtype syn true root (syms, pos) of |
888
3a1de9454d13
improved read_xrules: patterns no longer read twice;
wenzelm
parents:
882
diff
changeset
|
846 |
[ast] => constify ast |
27786
4525c6f5d753
added read_token -- with optional YXML encoding of position;
wenzelm
parents:
27768
diff
changeset
|
847 |
| _ => error (ambiguity_msg pos)) |
4525c6f5d753
added read_token -- with optional YXML encoding of position;
wenzelm
parents:
27768
diff
changeset
|
848 |
end; |
0 | 849 |
|
1158 | 850 |
fun prep_rules rd_pat raw_rules = |
4618 | 851 |
let val rules = map (map_trrule rd_pat) raw_rules in |
19482
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
wenzelm
parents:
19375
diff
changeset
|
852 |
(map check_rule (map_filter parse_rule rules), |
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
wenzelm
parents:
19375
diff
changeset
|
853 |
map check_rule (map_filter print_rule rules)) |
1158 | 854 |
end |
18 | 855 |
|
19262 | 856 |
in |
857 |
||
858 |
val cert_rules = prep_rules I; |
|
859 |
||
21772 | 860 |
fun read_rules ctxt is_logtype syn = |
861 |
prep_rules (read_pattern ctxt is_logtype syn); |
|
19262 | 862 |
|
863 |
end; |
|
864 |
||
18 | 865 |
|
866 |
||
24923 | 867 |
(** unparse terms, typs, sorts **) |
18 | 868 |
|
24923 | 869 |
local |
870 |
||
32784 | 871 |
fun unparse_t t_to_ast prt_t markup ctxt (Syntax (tabs, _)) curried t = |
0 | 872 |
let |
35429
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35412
diff
changeset
|
873 |
val {consts, print_trtab, print_ruletab, print_ast_trtab, tokentrtab, prtabs, ...} = tabs; |
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35412
diff
changeset
|
874 |
val ast = t_to_ast consts ctxt (lookup_tr' print_trtab) t; |
0 | 875 |
in |
23631 | 876 |
Pretty.markup markup (prt_t ctxt curried prtabs (lookup_tr' print_ast_trtab) |
24613 | 877 |
(lookup_tokentr tokentrtab (print_mode_value ())) |
23631 | 878 |
(Ast.normalize_ast (Symtab.lookup_list print_ruletab) ast)) |
0 | 879 |
end; |
880 |
||
24923 | 881 |
in |
882 |
||
35429
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35412
diff
changeset
|
883 |
fun standard_unparse_term idents extern = |
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35412
diff
changeset
|
884 |
unparse_t (Printer.term_to_ast idents) (Printer.pretty_term_ast extern) Markup.term; |
24923 | 885 |
|
35429
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35412
diff
changeset
|
886 |
fun standard_unparse_typ extern ctxt syn = |
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35412
diff
changeset
|
887 |
unparse_t (K Printer.typ_to_ast) (Printer.pretty_typ_ast extern) Markup.typ ctxt syn false; |
24923 | 888 |
|
35429
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35412
diff
changeset
|
889 |
fun standard_unparse_sort {extern_class} ctxt syn = |
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35412
diff
changeset
|
890 |
unparse_t (K Printer.sort_to_ast) |
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35412
diff
changeset
|
891 |
(Printer.pretty_typ_ast {extern_class = extern_class, extern_type = I}) |
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35412
diff
changeset
|
892 |
Markup.sort ctxt syn false; |
24923 | 893 |
|
894 |
end; |
|
0 | 895 |
|
896 |
||
897 |
||
19262 | 898 |
(** modify syntax **) |
383
fcea89074e4c
added incremental extension functions: extend_log_types, extend_type_gram,
wenzelm
parents:
330
diff
changeset
|
899 |
|
25394
db25c98f32e1
syntax operations: turned extend'' into update'' (absorb duplicates);
wenzelm
parents:
25387
diff
changeset
|
900 |
fun ext_syntax f decls = update_syntax mode_default (f decls); |
383
fcea89074e4c
added incremental extension functions: extend_log_types, extend_type_gram,
wenzelm
parents:
330
diff
changeset
|
901 |
|
37216
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
37146
diff
changeset
|
902 |
val update_trfuns = ext_syntax Syn_Ext.syn_ext_trfuns; |
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
37146
diff
changeset
|
903 |
val update_advanced_trfuns = ext_syntax Syn_Ext.syn_ext_advanced_trfuns; |
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
37146
diff
changeset
|
904 |
val extend_tokentrfuns = ext_syntax Syn_Ext.syn_ext_tokentrfuns; |
5692 | 905 |
|
35412
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
wenzelm
parents:
35394
diff
changeset
|
906 |
fun update_type_gram add prmode decls = |
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
wenzelm
parents:
35394
diff
changeset
|
907 |
(if add then update_syntax else remove_syntax) prmode (Mixfix.syn_ext_types decls); |
25387 | 908 |
|
35412
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
wenzelm
parents:
35394
diff
changeset
|
909 |
fun update_const_gram add is_logtype prmode decls = |
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
wenzelm
parents:
35394
diff
changeset
|
910 |
(if add then update_syntax else remove_syntax) prmode (Mixfix.syn_ext_consts is_logtype decls); |
15755
50ac97ebe7d8
expect translations functions to be stamped already;
wenzelm
parents:
15574
diff
changeset
|
911 |
|
25394
db25c98f32e1
syntax operations: turned extend'' into update'' (absorb duplicates);
wenzelm
parents:
25387
diff
changeset
|
912 |
fun update_trrules ctxt is_logtype syn = |
37216
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
37146
diff
changeset
|
913 |
ext_syntax Syn_Ext.syn_ext_rules o read_rules ctxt is_logtype syn; |
19262 | 914 |
|
21772 | 915 |
fun remove_trrules ctxt is_logtype syn = |
37216
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
37146
diff
changeset
|
916 |
remove_syntax mode_default o Syn_Ext.syn_ext_rules o read_rules ctxt is_logtype syn; |
19262 | 917 |
|
37216
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
37146
diff
changeset
|
918 |
val update_trrules_i = ext_syntax Syn_Ext.syn_ext_rules o cert_rules; |
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
37146
diff
changeset
|
919 |
val remove_trrules_i = remove_syntax mode_default o Syn_Ext.syn_ext_rules o cert_rules; |
5692 | 920 |
|
921 |
||
15833
78109c7012ed
removed token_trans.ML (some content moved to syn_ext.ML);
wenzelm
parents:
15759
diff
changeset
|
922 |
(*export parts of internal Syntax structures*) |
37216
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
37146
diff
changeset
|
923 |
open Lexicon Syn_Ext Ast Parser Type_Ext Syn_Trans Mixfix Printer; |
2366 | 924 |
|
0 | 925 |
end; |
5692 | 926 |
|
35130 | 927 |
structure Basic_Syntax: BASIC_SYNTAX = Syntax; |
928 |
open Basic_Syntax; |
|
23923 | 929 |
|
26684 | 930 |
forget_structure "Ast"; |
37216
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
37146
diff
changeset
|
931 |
forget_structure "Syn_Ext"; |
26684 | 932 |
forget_structure "Parser"; |
37216
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
37146
diff
changeset
|
933 |
forget_structure "Type_Ext"; |
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
37146
diff
changeset
|
934 |
forget_structure "Syn_Trans"; |
26684 | 935 |
forget_structure "Mixfix"; |
936 |
forget_structure "Printer"; |