author | wenzelm |
Sat, 07 Dec 2024 23:50:18 +0100 | |
changeset 81558 | b57996a0688c |
parent 81241 | 3b49bf00c8e4 |
permissions | -rw-r--r-- |
18 | 1 |
(* Title: Pure/Syntax/ast.ML |
0 | 2 |
Author: Markus Wenzel, TU Muenchen |
3 |
||
18 | 4 |
Abstract syntax trees, translation rules, matching and normalization of asts. |
0 | 5 |
*) |
6 |
||
42224
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
wenzelm
parents:
42048
diff
changeset
|
7 |
signature AST = |
16609 | 8 |
sig |
0 | 9 |
datatype ast = |
10 |
Constant of string | |
|
11 |
Variable of string | |
|
12 |
Appl of ast list |
|
42224
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
wenzelm
parents:
42048
diff
changeset
|
13 |
val mk_appl: ast -> ast list -> ast |
81210 | 14 |
val constrain: ast -> ast -> ast |
258 | 15 |
exception AST of string * ast list |
80881 | 16 |
val ast_ord: ast ord |
17 |
structure Table: TABLE |
|
18 |
structure Set: SET |
|
81153 | 19 |
val pretty_var: string -> Pretty.T |
18 | 20 |
val pretty_ast: ast -> Pretty.T |
258 | 21 |
val pretty_rule: ast * ast -> Pretty.T |
42264 | 22 |
val strip_positions: ast -> ast |
81152 | 23 |
val rule_index: ast * ast -> string |
42224
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
wenzelm
parents:
42048
diff
changeset
|
24 |
val rule_error: ast * ast -> string option |
10913 | 25 |
val fold_ast: string -> ast list -> ast |
26 |
val fold_ast_p: string -> ast list * ast -> ast |
|
27 |
val unfold_ast: string -> ast -> ast list |
|
28 |
val unfold_ast_p: string -> ast -> ast list * ast |
|
42277 | 29 |
val trace: bool Config.T |
43347 | 30 |
val stats: bool Config.T |
81208
893b056542e7
more accurate treatment of constraints: restrict permissive mode to output;
wenzelm
parents:
81207
diff
changeset
|
31 |
val normalize: Proof.context -> {permissive_constraints: bool} -> |
893b056542e7
more accurate treatment of constraints: restrict permissive mode to output;
wenzelm
parents:
81207
diff
changeset
|
32 |
(string -> (ast * ast) list) -> ast -> ast |
16609 | 33 |
end; |
0 | 34 |
|
42224
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
wenzelm
parents:
42048
diff
changeset
|
35 |
structure Ast: AST = |
0 | 36 |
struct |
37 |
||
18 | 38 |
(** abstract syntax trees **) |
0 | 39 |
|
40 |
(*asts come in two flavours: |
|
18 | 41 |
- ordinary asts representing terms and typs: Variables are (often) treated |
42 |
like Constants; |
|
0 | 43 |
- patterns used as lhs and rhs in rules: Variables are placeholders for |
44 |
proper asts*) |
|
45 |
||
46 |
datatype ast = |
|
18 | 47 |
Constant of string | (*"not", "_abs", "fun"*) |
48 |
Variable of string | (*x, ?x, 'a, ?'a*) |
|
49 |
Appl of ast list; (*(f x y z), ("fun" 'a 'b), ("_abs" x t)*) |
|
0 | 50 |
|
51 |
(*the list of subasts of an Appl node has to contain at least 2 elements, i.e. |
|
52 |
there are no empty asts or nullary applications; use mk_appl for convenience*) |
|
18 | 53 |
fun mk_appl f [] = f |
54 |
| mk_appl f args = Appl (f :: args); |
|
0 | 55 |
|
81210 | 56 |
fun constrain a b = Appl [Constant "_constrain", a, b]; |
57 |
||
0 | 58 |
(*exception for system errors involving asts*) |
59 |
exception AST of string * ast list; |
|
60 |
||
61 |
||
80881 | 62 |
(* order *) |
63 |
||
64 |
local |
|
65 |
||
66 |
fun cons_nr (Constant _) = 0 |
|
67 |
| cons_nr (Variable _) = 1 |
|
68 |
| cons_nr (Appl _) = 2; |
|
69 |
||
70 |
in |
|
71 |
||
72 |
fun ast_ord asts = |
|
73 |
if pointer_eq asts then EQUAL |
|
74 |
else |
|
75 |
(case asts of |
|
76 |
(Constant a, Constant b) => fast_string_ord (a, b) |
|
77 |
| (Variable a, Variable b) => fast_string_ord (a, b) |
|
78 |
| (Appl a, Appl b) => list_ord ast_ord (a, b) |
|
79 |
| _ => int_ord (apply2 cons_nr asts)); |
|
80 |
||
81 |
end; |
|
82 |
||
83 |
structure Table = Table(type key = ast val ord = ast_ord); |
|
84 |
structure Set = Set(Table.Key); |
|
85 |
||
86 |
||
81152 | 87 |
(* print asts in a LISP-like style *) |
18 | 88 |
|
81153 | 89 |
fun pretty_var x = |
90 |
(case Term_Position.decode x of |
|
81218 | 91 |
[] => Pretty.str x |
81558
b57996a0688c
clarified term positions and markup: syntax = true means this is via concrete syntax;
wenzelm
parents:
81241
diff
changeset
|
92 |
| ps => Term_Position.pretty (map #pos ps)); |
81153 | 93 |
|
14599 | 94 |
fun pretty_ast (Constant a) = Pretty.quote (Pretty.str a) |
81153 | 95 |
| pretty_ast (Variable x) = pretty_var x |
38328 | 96 |
| pretty_ast (Appl asts) = Pretty.enclose "(" ")" (Pretty.breaks (map pretty_ast asts)); |
18 | 97 |
|
98 |
fun pretty_rule (lhs, rhs) = |
|
67519 | 99 |
Pretty.block [pretty_ast lhs, Pretty.str " \<leadsto>", Pretty.brk 1, pretty_ast rhs]; |
18 | 100 |
|
80809
4a64fc4d1cde
clarified signature: type ML_Pretty.pretty coincides with PolyML.pretty;
wenzelm
parents:
74232
diff
changeset
|
101 |
val _ = ML_system_pp (fn _ => fn _ => Pretty.to_ML o pretty_ast); |
62663 | 102 |
|
18 | 103 |
|
42264 | 104 |
(* strip_positions *) |
105 |
||
106 |
fun strip_positions (Appl ((t as Constant c) :: u :: (v as Variable x) :: asts)) = |
|
81232 | 107 |
if member (op =) Term_Position.markers c andalso Term_Position.detect x |
42264 | 108 |
then mk_appl (strip_positions u) (map strip_positions asts) |
109 |
else Appl (map strip_positions (t :: u :: v :: asts)) |
|
110 |
| strip_positions (Appl asts) = Appl (map strip_positions asts) |
|
111 |
| strip_positions ast = ast; |
|
112 |
||
113 |
||
81152 | 114 |
(* translation rules *) |
0 | 115 |
|
81152 | 116 |
fun rule_index (Constant a, _: ast) = a |
117 |
| rule_index (Appl (Constant a :: _), _) = a |
|
118 |
| rule_index _ = ""; |
|
0 | 119 |
|
32784 | 120 |
fun rule_error (lhs, rhs) = |
0 | 121 |
let |
19486 | 122 |
fun add_vars (Constant _) = I |
123 |
| add_vars (Variable x) = cons x |
|
124 |
| add_vars (Appl asts) = fold add_vars asts; |
|
0 | 125 |
|
19486 | 126 |
val lvars = add_vars lhs []; |
127 |
val rvars = add_vars rhs []; |
|
0 | 128 |
in |
19486 | 129 |
if has_duplicates (op =) lvars then SOME "duplicate vars in lhs" |
33038 | 130 |
else if not (subset (op =) (rvars, lvars)) then SOME "rhs contains extra variables" |
15531 | 131 |
else NONE |
0 | 132 |
end; |
133 |
||
134 |
||
81152 | 135 |
(* ast translation utilities *) |
0 | 136 |
|
137 |
fun fold_ast _ [] = raise Match |
|
138 |
| fold_ast _ [y] = y |
|
139 |
| fold_ast c (x :: xs) = Appl [Constant c, x, fold_ast c xs]; |
|
140 |
||
19473 | 141 |
fun fold_ast_p c = uncurry (fold_rev (fn x => fn xs => Appl [Constant c, x, xs])); |
0 | 142 |
|
143 |
||
81166 | 144 |
fun unfold_ast c (y as Appl [Appl [Constant "_constrain", Constant c', _], x, xs]) = |
145 |
if c = c' then x :: unfold_ast c xs else [y] |
|
146 |
| unfold_ast c (y as Appl [Constant c', x, xs]) = |
|
16609 | 147 |
if c = c' then x :: unfold_ast c xs else [y] |
0 | 148 |
| unfold_ast _ y = [y]; |
149 |
||
81166 | 150 |
fun unfold_ast_p c (y as Appl [Appl [Constant "_constrain", Constant c', _], x, xs]) = |
81207 | 151 |
if c = c' then unfold_ast_p c xs |>> cons x else ([], y) |
81166 | 152 |
| unfold_ast_p c (y as Appl [Constant c', x, xs]) = |
81207 | 153 |
if c = c' then unfold_ast_p c xs |>> cons x else ([], y) |
0 | 154 |
| unfold_ast_p _ y = ([], y); |
155 |
||
156 |
||
16609 | 157 |
|
0 | 158 |
(** normalization of asts **) |
159 |
||
160 |
(* match *) |
|
161 |
||
81156 | 162 |
local |
163 |
||
164 |
exception NO_MATCH; |
|
0 | 165 |
|
81234
ae0ccabd0aab
more liberal ast matching, notably for case-translations in HOLCF that operate on logical consts rather than syntax consts;
wenzelm
parents:
81232
diff
changeset
|
166 |
fun const_match0 (Constant a) b = a = b |
ae0ccabd0aab
more liberal ast matching, notably for case-translations in HOLCF that operate on logical consts rather than syntax consts;
wenzelm
parents:
81232
diff
changeset
|
167 |
| const_match0 (Variable a) b = a = b |
ae0ccabd0aab
more liberal ast matching, notably for case-translations in HOLCF that operate on logical consts rather than syntax consts;
wenzelm
parents:
81232
diff
changeset
|
168 |
| const_match0 _ _ = false; |
ae0ccabd0aab
more liberal ast matching, notably for case-translations in HOLCF that operate on logical consts rather than syntax consts;
wenzelm
parents:
81232
diff
changeset
|
169 |
|
ae0ccabd0aab
more liberal ast matching, notably for case-translations in HOLCF that operate on logical consts rather than syntax consts;
wenzelm
parents:
81232
diff
changeset
|
170 |
fun const_match true (Appl [Constant "_constrain", ast, _]) b = const_match0 ast b |
ae0ccabd0aab
more liberal ast matching, notably for case-translations in HOLCF that operate on logical consts rather than syntax consts;
wenzelm
parents:
81232
diff
changeset
|
171 |
| const_match false (Appl [Constant "_constrain", ast, Variable x]) b = |
ae0ccabd0aab
more liberal ast matching, notably for case-translations in HOLCF that operate on logical consts rather than syntax consts;
wenzelm
parents:
81232
diff
changeset
|
172 |
Term_Position.detect x andalso const_match0 ast b |
ae0ccabd0aab
more liberal ast matching, notably for case-translations in HOLCF that operate on logical consts rather than syntax consts;
wenzelm
parents:
81232
diff
changeset
|
173 |
| const_match _ a b = const_match0 a b; |
81166 | 174 |
|
81208
893b056542e7
more accurate treatment of constraints: restrict permissive mode to output;
wenzelm
parents:
81207
diff
changeset
|
175 |
fun match1 p ast (Constant b) env = if const_match p ast b then env else raise NO_MATCH |
81234
ae0ccabd0aab
more liberal ast matching, notably for case-translations in HOLCF that operate on logical consts rather than syntax consts;
wenzelm
parents:
81232
diff
changeset
|
176 |
| match1 _ ast (Variable x) env = Symtab.update (x, ast) env |
81208
893b056542e7
more accurate treatment of constraints: restrict permissive mode to output;
wenzelm
parents:
81207
diff
changeset
|
177 |
| match1 p (Appl asts) (Appl pats) env = match2 p asts pats env |
81234
ae0ccabd0aab
more liberal ast matching, notably for case-translations in HOLCF that operate on logical consts rather than syntax consts;
wenzelm
parents:
81232
diff
changeset
|
178 |
| match1 _ _ _ _ = raise NO_MATCH |
81208
893b056542e7
more accurate treatment of constraints: restrict permissive mode to output;
wenzelm
parents:
81207
diff
changeset
|
179 |
and match2 p (ast :: asts) (pat :: pats) env = match1 p ast pat env |> match2 p asts pats |
893b056542e7
more accurate treatment of constraints: restrict permissive mode to output;
wenzelm
parents:
81207
diff
changeset
|
180 |
| match2 _ [] [] env = env |
893b056542e7
more accurate treatment of constraints: restrict permissive mode to output;
wenzelm
parents:
81207
diff
changeset
|
181 |
| match2 _ _ _ _ = raise NO_MATCH; |
81156 | 182 |
|
183 |
in |
|
184 |
||
81208
893b056542e7
more accurate treatment of constraints: restrict permissive mode to output;
wenzelm
parents:
81207
diff
changeset
|
185 |
fun match p ast pat = |
81151 | 186 |
let |
1127
42ec82147d83
changed macro expander such that patterns also match prefixes of appls;
wenzelm
parents:
922
diff
changeset
|
187 |
val (head, args) = |
81156 | 188 |
(case (ast, pat) of |
1127
42ec82147d83
changed macro expander such that patterns also match prefixes of appls;
wenzelm
parents:
922
diff
changeset
|
189 |
(Appl asts, Appl pats) => |
42ec82147d83
changed macro expander such that patterns also match prefixes of appls;
wenzelm
parents:
922
diff
changeset
|
190 |
let val a = length asts and p = length pats in |
33957 | 191 |
if a > p then (Appl (take p asts), drop p asts) |
81156 | 192 |
else (ast, []) |
1127
42ec82147d83
changed macro expander such that patterns also match prefixes of appls;
wenzelm
parents:
922
diff
changeset
|
193 |
end |
81156 | 194 |
| _ => (ast, [])); |
81208
893b056542e7
more accurate treatment of constraints: restrict permissive mode to output;
wenzelm
parents:
81207
diff
changeset
|
195 |
in SOME (Symtab.build (match1 p head pat), args) handle NO_MATCH => NONE end; |
81151 | 196 |
|
197 |
end; |
|
0 | 198 |
|
199 |
||
18 | 200 |
(* normalize *) |
0 | 201 |
|
69575 | 202 |
val trace = Config.declare_bool ("syntax_ast_trace", \<^here>) (K false); |
203 |
val stats = Config.declare_bool ("syntax_ast_stats", \<^here>) (K false); |
|
0 | 204 |
|
81151 | 205 |
local |
206 |
||
207 |
fun subst _ (ast as Constant _) = ast |
|
208 |
| subst env (Variable x) = the (Symtab.lookup env x) |
|
209 |
| subst env (Appl asts) = Appl (map (subst env) asts); |
|
210 |
||
81171 | 211 |
fun head_name (Constant a) = a |
212 |
| head_name (Variable a) = a |
|
213 |
| head_name (Appl [Constant "_constrain", Constant a, _]) = a |
|
214 |
| head_name (Appl (Appl [Constant "_constrain", Constant a, _] :: _)) = a |
|
215 |
| head_name (Appl (Constant a :: _)) = a |
|
216 |
| head_name (Appl (Variable a :: _)) = a |
|
217 |
| head_name _ = ""; |
|
81151 | 218 |
|
42048
afd11ca8e018
support for encoded positions (for id_position, longid_position) as pseudo type-constraints -- still inactive;
wenzelm
parents:
41377
diff
changeset
|
219 |
fun message head body = |
afd11ca8e018
support for encoded positions (for id_position, longid_position) as pseudo type-constraints -- still inactive;
wenzelm
parents:
41377
diff
changeset
|
220 |
Pretty.string_of (Pretty.block [Pretty.str head, Pretty.brk 1, body]); |
afd11ca8e018
support for encoded positions (for id_position, longid_position) as pseudo type-constraints -- still inactive;
wenzelm
parents:
41377
diff
changeset
|
221 |
|
81240 | 222 |
fun tracing_if false _ = () |
223 |
| tracing_if true msg = tracing (msg ()); |
|
224 |
||
81151 | 225 |
in |
226 |
||
41377
390c53904220
configuration option "syntax_ast_trace" and "syntax_ast_stat";
wenzelm
parents:
38328
diff
changeset
|
227 |
(*the normalizer works yoyo-like: top-down, bottom-up, top-down, ...*) |
81208
893b056542e7
more accurate treatment of constraints: restrict permissive mode to output;
wenzelm
parents:
81207
diff
changeset
|
228 |
fun normalize ctxt {permissive_constraints = p} get_rules pre_ast = |
0 | 229 |
let |
42277 | 230 |
val trace = Config.get ctxt trace; |
43347 | 231 |
val stats = Config.get ctxt stats; |
41377
390c53904220
configuration option "syntax_ast_trace" and "syntax_ast_stat";
wenzelm
parents:
38328
diff
changeset
|
232 |
|
81240 | 233 |
val matches_failed = Unsynchronized.ref 0; |
234 |
val changes = Unsynchronized.ref 0; |
|
32738 | 235 |
val passes = Unsynchronized.ref 0; |
0 | 236 |
|
81241 | 237 |
fun test_changes () = |
238 |
let val old_changes = ! changes |
|
239 |
in fn () => old_changes <> ! changes end; |
|
240 |
||
81240 | 241 |
fun rewrite1 ast (lhs, rhs) = |
242 |
(case match p ast lhs of |
|
243 |
NONE => (Unsynchronized.inc matches_failed; NONE) |
|
244 |
| SOME (env, args) => (Unsynchronized.inc changes; SOME (mk_appl (subst env rhs) args))); |
|
0 | 245 |
|
81171 | 246 |
fun rewrite2 a ast = |
81240 | 247 |
(case get_first (rewrite1 ast) (get_rules a) of |
81171 | 248 |
NONE => if a = "" then NONE else rewrite2 "" ast |
249 |
| some => some); |
|
0 | 250 |
|
81240 | 251 |
fun rewrote rule = tracing_if trace (fn () => message "rewrote:" (pretty_rule rule)); |
0 | 252 |
|
81151 | 253 |
fun norm1 ast = |
81240 | 254 |
(case rewrite2 (head_name ast) ast of |
255 |
SOME ast' => (rewrote (ast, ast'); norm1 ast') |
|
15531 | 256 |
| NONE => ast); |
0 | 257 |
|
81151 | 258 |
fun norm2 ast = |
259 |
(case norm1 ast of |
|
81240 | 260 |
Appl subs => |
0 | 261 |
let |
81241 | 262 |
val changed = test_changes (); |
81240 | 263 |
val ast' = Appl (map norm2 subs); |
81241 | 264 |
in if changed () then norm1 ast' else ast' end |
81240 | 265 |
| atomic => atomic); |
0 | 266 |
|
81151 | 267 |
fun norm ast = |
0 | 268 |
let |
81241 | 269 |
val changed = test_changes (); |
81240 | 270 |
val ast' = norm2 ast; |
271 |
val _ = Unsynchronized.inc passes; |
|
81241 | 272 |
in if changed () then norm ast' else ast' end; |
0 | 273 |
|
274 |
||
81240 | 275 |
val _ = tracing_if trace (fn () => message "pre:" (pretty_ast pre_ast)); |
81151 | 276 |
val post_ast = norm pre_ast; |
21962 | 277 |
val _ = |
81240 | 278 |
tracing_if (trace orelse stats) (fn () => |
279 |
message "post:" (pretty_ast post_ast) ^ "\nnormalize: " ^ |
|
280 |
string_of_int (! passes) ^ " passes, " ^ |
|
281 |
string_of_int (! changes) ^ " changes, " ^ |
|
282 |
string_of_int (! matches_failed) ^ " matches failed"); |
|
21962 | 283 |
in post_ast end; |
0 | 284 |
|
1506 | 285 |
end; |
81151 | 286 |
|
287 |
end; |