| author | wenzelm | 
| Tue, 23 Jan 2024 16:30:29 +0100 | |
| changeset 79519 | 557f00504bb6 | 
| parent 74232 | 1091880266e5 | 
| child 80809 | 4a64fc4d1cde | 
| 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  | 
| 258 | 14  | 
exception AST of string * ast list  | 
| 18 | 15  | 
val pretty_ast: ast -> Pretty.T  | 
| 258 | 16  | 
val pretty_rule: ast * ast -> Pretty.T  | 
| 42264 | 17  | 
val strip_positions: ast -> ast  | 
| 
42224
 
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
 
wenzelm 
parents: 
42048 
diff
changeset
 | 
18  | 
val head_of_rule: ast * ast -> string  | 
| 
 
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
 
wenzelm 
parents: 
42048 
diff
changeset
 | 
19  | 
val rule_error: ast * ast -> string option  | 
| 10913 | 20  | 
val fold_ast: string -> ast list -> ast  | 
21  | 
val fold_ast_p: string -> ast list * ast -> ast  | 
|
22  | 
val unfold_ast: string -> ast -> ast list  | 
|
23  | 
val unfold_ast_p: string -> ast -> ast list * ast  | 
|
| 42277 | 24  | 
val trace: bool Config.T  | 
| 43347 | 25  | 
val stats: bool Config.T  | 
| 
41377
 
390c53904220
configuration option "syntax_ast_trace" and "syntax_ast_stat";
 
wenzelm 
parents: 
38328 
diff
changeset
 | 
26  | 
val normalize: Proof.context -> (string -> (ast * ast) list) -> ast -> ast  | 
| 16609 | 27  | 
end;  | 
| 0 | 28  | 
|
| 
42224
 
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
 
wenzelm 
parents: 
42048 
diff
changeset
 | 
29  | 
structure Ast: AST =  | 
| 0 | 30  | 
struct  | 
31  | 
||
| 18 | 32  | 
(** abstract syntax trees **)  | 
| 0 | 33  | 
|
34  | 
(*asts come in two flavours:  | 
|
| 18 | 35  | 
- ordinary asts representing terms and typs: Variables are (often) treated  | 
36  | 
like Constants;  | 
|
| 0 | 37  | 
- patterns used as lhs and rhs in rules: Variables are placeholders for  | 
38  | 
proper asts*)  | 
|
39  | 
||
40  | 
datatype ast =  | 
|
| 18 | 41  | 
Constant of string | (*"not", "_abs", "fun"*)  | 
42  | 
Variable of string | (*x, ?x, 'a, ?'a*)  | 
|
43  | 
  Appl of ast list;       (*(f x y z), ("fun" 'a 'b), ("_abs" x t)*)
 | 
|
| 0 | 44  | 
|
45  | 
(*the list of subasts of an Appl node has to contain at least 2 elements, i.e.  | 
|
46  | 
there are no empty asts or nullary applications; use mk_appl for convenience*)  | 
|
| 18 | 47  | 
fun mk_appl f [] = f  | 
48  | 
| mk_appl f args = Appl (f :: args);  | 
|
| 0 | 49  | 
|
50  | 
(*exception for system errors involving asts*)  | 
|
51  | 
exception AST of string * ast list;  | 
|
52  | 
||
53  | 
||
54  | 
||
| 18 | 55  | 
(** print asts in a LISP-like style **)  | 
56  | 
||
| 14599 | 57  | 
fun pretty_ast (Constant a) = Pretty.quote (Pretty.str a)  | 
| 
42048
 
afd11ca8e018
support for encoded positions (for id_position, longid_position) as pseudo type-constraints -- still inactive;
 
wenzelm 
parents: 
41377 
diff
changeset
 | 
58  | 
| pretty_ast (Variable x) =  | 
| 42264 | 59  | 
(case Term_Position.decode x of  | 
60  | 
SOME pos => Term_Position.pretty pos  | 
|
| 
42048
 
afd11ca8e018
support for encoded positions (for id_position, longid_position) as pseudo type-constraints -- still inactive;
 
wenzelm 
parents: 
41377 
diff
changeset
 | 
61  | 
| NONE => Pretty.str x)  | 
| 38328 | 62  | 
  | pretty_ast (Appl asts) = Pretty.enclose "(" ")" (Pretty.breaks (map pretty_ast asts));
 | 
| 18 | 63  | 
|
64  | 
fun pretty_rule (lhs, rhs) =  | 
|
| 67519 | 65  | 
Pretty.block [pretty_ast lhs, Pretty.str " \<leadsto>", Pretty.brk 1, pretty_ast rhs];  | 
| 18 | 66  | 
|
| 
62819
 
d3ff367a16a0
careful export of type-dependent functions, without losing their special status;
 
wenzelm 
parents: 
62663 
diff
changeset
 | 
67  | 
val _ = ML_system_pp (fn _ => fn _ => Pretty.to_polyml o pretty_ast);  | 
| 62663 | 68  | 
|
| 18 | 69  | 
|
| 42264 | 70  | 
(* strip_positions *)  | 
71  | 
||
72  | 
fun strip_positions (Appl ((t as Constant c) :: u :: (v as Variable x) :: asts)) =  | 
|
| 
52175
 
626a757d3c2d
uniform Term_Position.markers (cf. dbadb4d03cbc);
 
wenzelm 
parents: 
43347 
diff
changeset
 | 
73  | 
if member (op =) Term_Position.markers c andalso is_some (Term_Position.decode x)  | 
| 42264 | 74  | 
then mk_appl (strip_positions u) (map strip_positions asts)  | 
75  | 
else Appl (map strip_positions (t :: u :: v :: asts))  | 
|
76  | 
| strip_positions (Appl asts) = Appl (map strip_positions asts)  | 
|
77  | 
| strip_positions ast = ast;  | 
|
78  | 
||
79  | 
||
80  | 
(* head_of_ast and head_of_rule *)  | 
|
| 0 | 81  | 
|
| 
9372
 
7834e56e2277
AST translation rules no longer require constant head on LHS;
 
wenzelm 
parents: 
8997 
diff
changeset
 | 
82  | 
fun head_of_ast (Constant a) = a  | 
| 
 
7834e56e2277
AST translation rules no longer require constant head on LHS;
 
wenzelm 
parents: 
8997 
diff
changeset
 | 
83  | 
| head_of_ast (Appl (Constant a :: _)) = a  | 
| 
 
7834e56e2277
AST translation rules no longer require constant head on LHS;
 
wenzelm 
parents: 
8997 
diff
changeset
 | 
84  | 
| head_of_ast _ = "";  | 
| 0 | 85  | 
|
| 
9372
 
7834e56e2277
AST translation rules no longer require constant head on LHS;
 
wenzelm 
parents: 
8997 
diff
changeset
 | 
86  | 
fun head_of_rule (lhs, _) = head_of_ast lhs;  | 
| 0 | 87  | 
|
88  | 
||
89  | 
||
| 18 | 90  | 
(** check translation rules **)  | 
| 0 | 91  | 
|
| 32784 | 92  | 
fun rule_error (lhs, rhs) =  | 
| 0 | 93  | 
let  | 
| 19486 | 94  | 
fun add_vars (Constant _) = I  | 
95  | 
| add_vars (Variable x) = cons x  | 
|
96  | 
| add_vars (Appl asts) = fold add_vars asts;  | 
|
| 0 | 97  | 
|
| 19486 | 98  | 
val lvars = add_vars lhs [];  | 
99  | 
val rvars = add_vars rhs [];  | 
|
| 0 | 100  | 
in  | 
| 19486 | 101  | 
if has_duplicates (op =) lvars then SOME "duplicate vars in lhs"  | 
| 33038 | 102  | 
else if not (subset (op =) (rvars, lvars)) then SOME "rhs contains extra variables"  | 
| 15531 | 103  | 
else NONE  | 
| 0 | 104  | 
end;  | 
105  | 
||
106  | 
||
107  | 
||
| 18 | 108  | 
(** ast translation utilities **)  | 
| 0 | 109  | 
|
110  | 
(* fold asts *)  | 
|
111  | 
||
112  | 
fun fold_ast _ [] = raise Match  | 
|
113  | 
| fold_ast _ [y] = y  | 
|
114  | 
| fold_ast c (x :: xs) = Appl [Constant c, x, fold_ast c xs];  | 
|
115  | 
||
| 19473 | 116  | 
fun fold_ast_p c = uncurry (fold_rev (fn x => fn xs => Appl [Constant c, x, xs]));  | 
| 0 | 117  | 
|
118  | 
||
119  | 
(* unfold asts *)  | 
|
120  | 
||
121  | 
fun unfold_ast c (y as Appl [Constant c', x, xs]) =  | 
|
| 16609 | 122  | 
if c = c' then x :: unfold_ast c xs else [y]  | 
| 0 | 123  | 
| unfold_ast _ y = [y];  | 
124  | 
||
125  | 
fun unfold_ast_p c (y as Appl [Constant c', x, xs]) =  | 
|
| 18 | 126  | 
if c = c' then apfst (cons x) (unfold_ast_p c xs)  | 
| 0 | 127  | 
else ([], y)  | 
128  | 
| unfold_ast_p _ y = ([], y);  | 
|
129  | 
||
130  | 
||
| 16609 | 131  | 
|
| 0 | 132  | 
(** normalization of asts **)  | 
133  | 
||
134  | 
(* match *)  | 
|
135  | 
||
136  | 
fun match ast pat =  | 
|
137  | 
let  | 
|
138  | 
exception NO_MATCH;  | 
|
139  | 
||
| 
1127
 
42ec82147d83
changed macro expander such that patterns also match prefixes of appls;
 
wenzelm 
parents: 
922 
diff
changeset
 | 
140  | 
fun mtch (Constant a) (Constant b) env =  | 
| 0 | 141  | 
if a = b then env else raise NO_MATCH  | 
| 
1127
 
42ec82147d83
changed macro expander such that patterns also match prefixes of appls;
 
wenzelm 
parents: 
922 
diff
changeset
 | 
142  | 
| mtch (Variable a) (Constant b) env =  | 
| 0 | 143  | 
if a = b then env else raise NO_MATCH  | 
| 17412 | 144  | 
| mtch ast (Variable x) env = Symtab.update (x, ast) env  | 
| 
1127
 
42ec82147d83
changed macro expander such that patterns also match prefixes of appls;
 
wenzelm 
parents: 
922 
diff
changeset
 | 
145  | 
| mtch (Appl asts) (Appl pats) env = mtch_lst asts pats env  | 
| 
 
42ec82147d83
changed macro expander such that patterns also match prefixes of appls;
 
wenzelm 
parents: 
922 
diff
changeset
 | 
146  | 
| mtch _ _ _ = raise NO_MATCH  | 
| 
 
42ec82147d83
changed macro expander such that patterns also match prefixes of appls;
 
wenzelm 
parents: 
922 
diff
changeset
 | 
147  | 
and mtch_lst (ast :: asts) (pat :: pats) env =  | 
| 
 
42ec82147d83
changed macro expander such that patterns also match prefixes of appls;
 
wenzelm 
parents: 
922 
diff
changeset
 | 
148  | 
mtch_lst asts pats (mtch ast pat env)  | 
| 
 
42ec82147d83
changed macro expander such that patterns also match prefixes of appls;
 
wenzelm 
parents: 
922 
diff
changeset
 | 
149  | 
| mtch_lst [] [] env = env  | 
| 
 
42ec82147d83
changed macro expander such that patterns also match prefixes of appls;
 
wenzelm 
parents: 
922 
diff
changeset
 | 
150  | 
| mtch_lst _ _ _ = raise NO_MATCH;  | 
| 
 
42ec82147d83
changed macro expander such that patterns also match prefixes of appls;
 
wenzelm 
parents: 
922 
diff
changeset
 | 
151  | 
|
| 
 
42ec82147d83
changed macro expander such that patterns also match prefixes of appls;
 
wenzelm 
parents: 
922 
diff
changeset
 | 
152  | 
val (head, args) =  | 
| 
 
42ec82147d83
changed macro expander such that patterns also match prefixes of appls;
 
wenzelm 
parents: 
922 
diff
changeset
 | 
153  | 
(case (ast, pat) of  | 
| 
 
42ec82147d83
changed macro expander such that patterns also match prefixes of appls;
 
wenzelm 
parents: 
922 
diff
changeset
 | 
154  | 
(Appl asts, Appl pats) =>  | 
| 
 
42ec82147d83
changed macro expander such that patterns also match prefixes of appls;
 
wenzelm 
parents: 
922 
diff
changeset
 | 
155  | 
let val a = length asts and p = length pats in  | 
| 33957 | 156  | 
if a > p then (Appl (take p asts), drop p asts)  | 
| 
1127
 
42ec82147d83
changed macro expander such that patterns also match prefixes of appls;
 
wenzelm 
parents: 
922 
diff
changeset
 | 
157  | 
else (ast, [])  | 
| 
 
42ec82147d83
changed macro expander such that patterns also match prefixes of appls;
 
wenzelm 
parents: 
922 
diff
changeset
 | 
158  | 
end  | 
| 
 
42ec82147d83
changed macro expander such that patterns also match prefixes of appls;
 
wenzelm 
parents: 
922 
diff
changeset
 | 
159  | 
| _ => (ast, []));  | 
| 0 | 160  | 
in  | 
| 74232 | 161  | 
SOME (Symtab.build (mtch head pat), args) handle NO_MATCH => NONE  | 
| 0 | 162  | 
end;  | 
163  | 
||
164  | 
||
| 18 | 165  | 
(* normalize *)  | 
| 0 | 166  | 
|
| 69575 | 167  | 
val trace = Config.declare_bool ("syntax_ast_trace", \<^here>) (K false);
 | 
168  | 
val stats = Config.declare_bool ("syntax_ast_stats", \<^here>) (K false);
 | 
|
| 0 | 169  | 
|
| 
42048
 
afd11ca8e018
support for encoded positions (for id_position, longid_position) as pseudo type-constraints -- still inactive;
 
wenzelm 
parents: 
41377 
diff
changeset
 | 
170  | 
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
 | 
171  | 
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
 | 
172  | 
|
| 
41377
 
390c53904220
configuration option "syntax_ast_trace" and "syntax_ast_stat";
 
wenzelm 
parents: 
38328 
diff
changeset
 | 
173  | 
(*the normalizer works yoyo-like: top-down, bottom-up, top-down, ...*)  | 
| 
 
390c53904220
configuration option "syntax_ast_trace" and "syntax_ast_stat";
 
wenzelm 
parents: 
38328 
diff
changeset
 | 
174  | 
fun normalize ctxt get_rules pre_ast =  | 
| 0 | 175  | 
let  | 
| 42277 | 176  | 
val trace = Config.get ctxt trace;  | 
| 43347 | 177  | 
val stats = Config.get ctxt stats;  | 
| 
41377
 
390c53904220
configuration option "syntax_ast_trace" and "syntax_ast_stat";
 
wenzelm 
parents: 
38328 
diff
changeset
 | 
178  | 
|
| 32738 | 179  | 
val passes = Unsynchronized.ref 0;  | 
180  | 
val failed_matches = Unsynchronized.ref 0;  | 
|
181  | 
val changes = Unsynchronized.ref 0;  | 
|
| 0 | 182  | 
|
| 18 | 183  | 
fun subst _ (ast as Constant _) = ast  | 
| 17412 | 184  | 
| subst env (Variable x) = the (Symtab.lookup env x)  | 
| 0 | 185  | 
| subst env (Appl asts) = Appl (map (subst env) asts);  | 
186  | 
||
| 11733 | 187  | 
fun try_rules ((lhs, rhs) :: pats) ast =  | 
| 0 | 188  | 
(case match ast lhs of  | 
| 15531 | 189  | 
SOME (env, args) =>  | 
| 32738 | 190  | 
(Unsynchronized.inc changes; SOME (mk_appl (subst env rhs) args))  | 
191  | 
| NONE => (Unsynchronized.inc failed_matches; try_rules pats ast))  | 
|
| 15531 | 192  | 
| try_rules [] _ = NONE;  | 
| 11733 | 193  | 
val try_headless_rules = try_rules (get_rules "");  | 
| 0 | 194  | 
|
| 11733 | 195  | 
fun try ast a =  | 
196  | 
(case try_rules (get_rules a) ast of  | 
|
| 15531 | 197  | 
NONE => try_headless_rules ast  | 
| 11733 | 198  | 
| some => some);  | 
| 0 | 199  | 
|
200  | 
fun rewrite (ast as Constant a) = try ast a  | 
|
201  | 
| rewrite (ast as Variable a) = try ast a  | 
|
202  | 
| rewrite (ast as Appl (Constant a :: _)) = try ast a  | 
|
203  | 
| rewrite (ast as Appl (Variable a :: _)) = try ast a  | 
|
| 11733 | 204  | 
| rewrite ast = try_headless_rules ast;  | 
| 0 | 205  | 
|
206  | 
fun rewrote old_ast new_ast =  | 
|
| 
42048
 
afd11ca8e018
support for encoded positions (for id_position, longid_position) as pseudo type-constraints -- still inactive;
 
wenzelm 
parents: 
41377 
diff
changeset
 | 
207  | 
if trace then tracing (message "rewrote:" (pretty_rule (old_ast, new_ast)))  | 
| 21962 | 208  | 
else ();  | 
| 0 | 209  | 
|
210  | 
fun norm_root ast =  | 
|
211  | 
(case rewrite ast of  | 
|
| 15531 | 212  | 
SOME new_ast => (rewrote ast new_ast; norm_root new_ast)  | 
213  | 
| NONE => ast);  | 
|
| 0 | 214  | 
|
215  | 
fun norm ast =  | 
|
216  | 
(case norm_root ast of  | 
|
217  | 
Appl sub_asts =>  | 
|
218  | 
let  | 
|
219  | 
val old_changes = ! changes;  | 
|
220  | 
val new_ast = Appl (map norm sub_asts);  | 
|
221  | 
in  | 
|
222  | 
if old_changes = ! changes then new_ast else norm_root new_ast  | 
|
223  | 
end  | 
|
224  | 
| atomic_ast => atomic_ast);  | 
|
225  | 
||
226  | 
fun normal ast =  | 
|
227  | 
let  | 
|
228  | 
val old_changes = ! changes;  | 
|
229  | 
val new_ast = norm ast;  | 
|
230  | 
in  | 
|
| 32738 | 231  | 
Unsynchronized.inc passes;  | 
| 0 | 232  | 
if old_changes = ! changes then new_ast else normal new_ast  | 
233  | 
end;  | 
|
234  | 
||
235  | 
||
| 
42048
 
afd11ca8e018
support for encoded positions (for id_position, longid_position) as pseudo type-constraints -- still inactive;
 
wenzelm 
parents: 
41377 
diff
changeset
 | 
236  | 
val _ = if trace then tracing (message "pre:" (pretty_ast pre_ast)) else ();  | 
| 
9372
 
7834e56e2277
AST translation rules no longer require constant head on LHS;
 
wenzelm 
parents: 
8997 
diff
changeset
 | 
237  | 
val post_ast = normal pre_ast;  | 
| 21962 | 238  | 
val _ =  | 
| 43347 | 239  | 
if trace orelse stats then  | 
| 
42048
 
afd11ca8e018
support for encoded positions (for id_position, longid_position) as pseudo type-constraints -- still inactive;
 
wenzelm 
parents: 
41377 
diff
changeset
 | 
240  | 
tracing (message "post:" (pretty_ast post_ast) ^ "\nnormalize: " ^  | 
| 21962 | 241  | 
string_of_int (! passes) ^ " passes, " ^  | 
242  | 
string_of_int (! changes) ^ " changes, " ^  | 
|
243  | 
string_of_int (! failed_matches) ^ " matches failed")  | 
|
244  | 
else ();  | 
|
245  | 
in post_ast end;  | 
|
| 0 | 246  | 
|
| 1506 | 247  | 
end;  |