| author | wenzelm | 
| Sat, 07 Jul 2007 00:14:56 +0200 | |
| changeset 23615 | 40ab945ef5ff | 
| parent 21962 | 279b129498b6 | 
| child 29565 | 3f8b24fcfbd6 | 
| permissions | -rw-r--r-- | 
| 18 | 1  | 
(* Title: Pure/Syntax/ast.ML  | 
| 0 | 2  | 
ID: $Id$  | 
3  | 
Author: Markus Wenzel, TU Muenchen  | 
|
4  | 
||
| 18 | 5  | 
Abstract syntax trees, translation rules, matching and normalization of asts.  | 
| 0 | 6  | 
*)  | 
7  | 
||
| 18 | 8  | 
signature AST0 =  | 
| 16609 | 9  | 
sig  | 
| 0 | 10  | 
datatype ast =  | 
11  | 
Constant of string |  | 
|
12  | 
Variable of string |  | 
|
13  | 
Appl of ast list  | 
|
| 258 | 14  | 
exception AST of string * ast list  | 
| 16609 | 15  | 
end;  | 
| 258 | 16  | 
|
17  | 
signature AST1 =  | 
|
| 16609 | 18  | 
sig  | 
| 258 | 19  | 
include AST0  | 
| 0 | 20  | 
val mk_appl: ast -> ast list -> ast  | 
| 18 | 21  | 
val str_of_ast: ast -> string  | 
22  | 
val pretty_ast: ast -> Pretty.T  | 
|
| 258 | 23  | 
val pretty_rule: ast * ast -> Pretty.T  | 
| 18 | 24  | 
val pprint_ast: ast -> pprint_args -> unit  | 
| 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  | 
|
| 8997 | 29  | 
val trace_ast: bool ref  | 
30  | 
val stat_ast: bool ref  | 
|
| 16609 | 31  | 
end;  | 
| 18 | 32  | 
|
33  | 
signature AST =  | 
|
| 16609 | 34  | 
sig  | 
| 258 | 35  | 
include AST1  | 
| 0 | 36  | 
val head_of_rule: ast * ast -> string  | 
37  | 
val rule_error: ast * ast -> string option  | 
|
| 
9372
 
7834e56e2277
AST translation rules no longer require constant head on LHS;
 
wenzelm 
parents: 
8997 
diff
changeset
 | 
38  | 
val normalize: bool -> bool -> (string -> (ast * ast) list) -> ast -> ast  | 
| 
 
7834e56e2277
AST translation rules no longer require constant head on LHS;
 
wenzelm 
parents: 
8997 
diff
changeset
 | 
39  | 
val normalize_ast: (string -> (ast * ast) list) -> ast -> ast  | 
| 16609 | 40  | 
end;  | 
| 0 | 41  | 
|
| 1506 | 42  | 
structure Ast : AST =  | 
| 0 | 43  | 
struct  | 
44  | 
||
| 18 | 45  | 
(** abstract syntax trees **)  | 
| 0 | 46  | 
|
47  | 
(*asts come in two flavours:  | 
|
| 18 | 48  | 
- ordinary asts representing terms and typs: Variables are (often) treated  | 
49  | 
like Constants;  | 
|
| 0 | 50  | 
- patterns used as lhs and rhs in rules: Variables are placeholders for  | 
51  | 
proper asts*)  | 
|
52  | 
||
53  | 
datatype ast =  | 
|
| 18 | 54  | 
Constant of string | (*"not", "_abs", "fun"*)  | 
55  | 
Variable of string | (*x, ?x, 'a, ?'a*)  | 
|
56  | 
  Appl of ast list;       (*(f x y z), ("fun" 'a 'b), ("_abs" x t)*)
 | 
|
| 0 | 57  | 
|
58  | 
||
59  | 
(*the list of subasts of an Appl node has to contain at least 2 elements, i.e.  | 
|
60  | 
there are no empty asts or nullary applications; use mk_appl for convenience*)  | 
|
61  | 
||
| 18 | 62  | 
fun mk_appl f [] = f  | 
63  | 
| mk_appl f args = Appl (f :: args);  | 
|
| 0 | 64  | 
|
65  | 
||
66  | 
(*exception for system errors involving asts*)  | 
|
67  | 
||
68  | 
exception AST of string * ast list;  | 
|
69  | 
||
70  | 
||
71  | 
||
| 18 | 72  | 
(** print asts in a LISP-like style **)  | 
73  | 
||
| 0 | 74  | 
fun str_of_ast (Constant a) = quote a  | 
75  | 
| str_of_ast (Variable x) = x  | 
|
76  | 
  | str_of_ast (Appl asts) = "(" ^ (space_implode " " (map str_of_ast asts)) ^ ")";
 | 
|
77  | 
||
| 14599 | 78  | 
fun pretty_ast (Constant a) = Pretty.quote (Pretty.str a)  | 
| 18 | 79  | 
| pretty_ast (Variable x) = Pretty.str x  | 
80  | 
| pretty_ast (Appl asts) =  | 
|
| 513 | 81  | 
      Pretty.enclose "(" ")" (Pretty.breaks (map pretty_ast asts));
 | 
| 18 | 82  | 
|
83  | 
val pprint_ast = Pretty.pprint o pretty_ast;  | 
|
84  | 
||
85  | 
fun pretty_rule (lhs, rhs) =  | 
|
| 235 | 86  | 
Pretty.block [pretty_ast lhs, Pretty.str " ->", Pretty.brk 2, pretty_ast rhs];  | 
| 18 | 87  | 
|
88  | 
||
| 0 | 89  | 
(* head_of_ast, head_of_rule *)  | 
90  | 
||
| 
9372
 
7834e56e2277
AST translation rules no longer require constant head on LHS;
 
wenzelm 
parents: 
8997 
diff
changeset
 | 
91  | 
fun head_of_ast (Constant a) = a  | 
| 
 
7834e56e2277
AST translation rules no longer require constant head on LHS;
 
wenzelm 
parents: 
8997 
diff
changeset
 | 
92  | 
| head_of_ast (Appl (Constant a :: _)) = a  | 
| 
 
7834e56e2277
AST translation rules no longer require constant head on LHS;
 
wenzelm 
parents: 
8997 
diff
changeset
 | 
93  | 
| head_of_ast _ = "";  | 
| 0 | 94  | 
|
| 
9372
 
7834e56e2277
AST translation rules no longer require constant head on LHS;
 
wenzelm 
parents: 
8997 
diff
changeset
 | 
95  | 
fun head_of_rule (lhs, _) = head_of_ast lhs;  | 
| 0 | 96  | 
|
97  | 
||
98  | 
||
| 18 | 99  | 
(** check translation rules **)  | 
| 0 | 100  | 
|
101  | 
fun rule_error (rule as (lhs, rhs)) =  | 
|
102  | 
let  | 
|
| 19486 | 103  | 
fun add_vars (Constant _) = I  | 
104  | 
| add_vars (Variable x) = cons x  | 
|
105  | 
| add_vars (Appl asts) = fold add_vars asts;  | 
|
| 0 | 106  | 
|
| 19486 | 107  | 
val lvars = add_vars lhs [];  | 
108  | 
val rvars = add_vars rhs [];  | 
|
| 0 | 109  | 
in  | 
| 19486 | 110  | 
if has_duplicates (op =) lvars then SOME "duplicate vars in lhs"  | 
| 15531 | 111  | 
else if not (rvars subset lvars) then SOME "rhs contains extra variables"  | 
112  | 
else NONE  | 
|
| 0 | 113  | 
end;  | 
114  | 
||
115  | 
||
116  | 
||
| 18 | 117  | 
(** ast translation utilities **)  | 
| 0 | 118  | 
|
119  | 
(* fold asts *)  | 
|
120  | 
||
121  | 
fun fold_ast _ [] = raise Match  | 
|
122  | 
| fold_ast _ [y] = y  | 
|
123  | 
| fold_ast c (x :: xs) = Appl [Constant c, x, fold_ast c xs];  | 
|
124  | 
||
| 19473 | 125  | 
fun fold_ast_p c = uncurry (fold_rev (fn x => fn xs => Appl [Constant c, x, xs]));  | 
| 0 | 126  | 
|
127  | 
||
128  | 
(* unfold asts *)  | 
|
129  | 
||
130  | 
fun unfold_ast c (y as Appl [Constant c', x, xs]) =  | 
|
| 16609 | 131  | 
if c = c' then x :: unfold_ast c xs else [y]  | 
| 0 | 132  | 
| unfold_ast _ y = [y];  | 
133  | 
||
134  | 
fun unfold_ast_p c (y as Appl [Constant c', x, xs]) =  | 
|
| 18 | 135  | 
if c = c' then apfst (cons x) (unfold_ast_p c xs)  | 
| 0 | 136  | 
else ([], y)  | 
137  | 
| unfold_ast_p _ y = ([], y);  | 
|
138  | 
||
139  | 
||
| 16609 | 140  | 
|
| 0 | 141  | 
(** normalization of asts **)  | 
142  | 
||
143  | 
(* match *)  | 
|
144  | 
||
145  | 
fun match ast pat =  | 
|
146  | 
let  | 
|
147  | 
exception NO_MATCH;  | 
|
148  | 
||
| 
1127
 
42ec82147d83
changed macro expander such that patterns also match prefixes of appls;
 
wenzelm 
parents: 
922 
diff
changeset
 | 
149  | 
fun mtch (Constant a) (Constant b) env =  | 
| 0 | 150  | 
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
 | 
151  | 
| mtch (Variable a) (Constant b) env =  | 
| 0 | 152  | 
if a = b then env else raise NO_MATCH  | 
| 17412 | 153  | 
| 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
 | 
154  | 
| 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
 | 
155  | 
| mtch _ _ _ = raise NO_MATCH  | 
| 
 
42ec82147d83
changed macro expander such that patterns also match prefixes of appls;
 
wenzelm 
parents: 
922 
diff
changeset
 | 
156  | 
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
 | 
157  | 
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
 | 
158  | 
| mtch_lst [] [] env = env  | 
| 
 
42ec82147d83
changed macro expander such that patterns also match prefixes of appls;
 
wenzelm 
parents: 
922 
diff
changeset
 | 
159  | 
| mtch_lst _ _ _ = raise NO_MATCH;  | 
| 
 
42ec82147d83
changed macro expander such that patterns also match prefixes of appls;
 
wenzelm 
parents: 
922 
diff
changeset
 | 
160  | 
|
| 
 
42ec82147d83
changed macro expander such that patterns also match prefixes of appls;
 
wenzelm 
parents: 
922 
diff
changeset
 | 
161  | 
val (head, args) =  | 
| 
 
42ec82147d83
changed macro expander such that patterns also match prefixes of appls;
 
wenzelm 
parents: 
922 
diff
changeset
 | 
162  | 
(case (ast, pat) of  | 
| 
 
42ec82147d83
changed macro expander such that patterns also match prefixes of appls;
 
wenzelm 
parents: 
922 
diff
changeset
 | 
163  | 
(Appl asts, Appl pats) =>  | 
| 
 
42ec82147d83
changed macro expander such that patterns also match prefixes of appls;
 
wenzelm 
parents: 
922 
diff
changeset
 | 
164  | 
let val a = length asts and p = length pats in  | 
| 15570 | 165  | 
if a > p then (Appl (Library.take (p, asts)), Library.drop (p, asts))  | 
| 
1127
 
42ec82147d83
changed macro expander such that patterns also match prefixes of appls;
 
wenzelm 
parents: 
922 
diff
changeset
 | 
166  | 
else (ast, [])  | 
| 
 
42ec82147d83
changed macro expander such that patterns also match prefixes of appls;
 
wenzelm 
parents: 
922 
diff
changeset
 | 
167  | 
end  | 
| 
 
42ec82147d83
changed macro expander such that patterns also match prefixes of appls;
 
wenzelm 
parents: 
922 
diff
changeset
 | 
168  | 
| _ => (ast, []));  | 
| 0 | 169  | 
in  | 
| 15531 | 170  | 
SOME (mtch head pat Symtab.empty, args) handle NO_MATCH => NONE  | 
| 0 | 171  | 
end;  | 
172  | 
||
173  | 
||
| 18 | 174  | 
(* normalize *)  | 
| 0 | 175  | 
|
176  | 
(*the normalizer works yoyo-like: top-down, bottom-up, top-down, ...*)  | 
|
177  | 
||
| 18 | 178  | 
fun normalize trace stat get_rules pre_ast =  | 
| 0 | 179  | 
let  | 
180  | 
val passes = ref 0;  | 
|
181  | 
val failed_matches = ref 0;  | 
|
182  | 
val changes = ref 0;  | 
|
183  | 
||
| 18 | 184  | 
fun subst _ (ast as Constant _) = ast  | 
| 17412 | 185  | 
| subst env (Variable x) = the (Symtab.lookup env x)  | 
| 0 | 186  | 
| subst env (Appl asts) = Appl (map (subst env) asts);  | 
187  | 
||
| 11733 | 188  | 
fun try_rules ((lhs, rhs) :: pats) ast =  | 
| 0 | 189  | 
(case match ast lhs of  | 
| 15531 | 190  | 
SOME (env, args) =>  | 
191  | 
(inc changes; SOME (mk_appl (subst env rhs) args))  | 
|
192  | 
| NONE => (inc failed_matches; try_rules pats ast))  | 
|
193  | 
| try_rules [] _ = NONE;  | 
|
| 11733 | 194  | 
val try_headless_rules = try_rules (get_rules "");  | 
| 0 | 195  | 
|
| 11733 | 196  | 
fun try ast a =  | 
197  | 
(case try_rules (get_rules a) ast of  | 
|
| 15531 | 198  | 
NONE => try_headless_rules ast  | 
| 11733 | 199  | 
| some => some);  | 
| 0 | 200  | 
|
201  | 
fun rewrite (ast as Constant a) = try ast a  | 
|
202  | 
| rewrite (ast as Variable a) = try ast a  | 
|
203  | 
| rewrite (ast as Appl (Constant a :: _)) = try ast a  | 
|
204  | 
| rewrite (ast as Appl (Variable a :: _)) = try ast a  | 
|
| 11733 | 205  | 
| rewrite ast = try_headless_rules ast;  | 
| 0 | 206  | 
|
207  | 
fun rewrote old_ast new_ast =  | 
|
| 21962 | 208  | 
if trace then  | 
209  | 
        tracing ("rewrote: " ^ str_of_ast old_ast ^ "  ->  " ^ str_of_ast new_ast)
 | 
|
210  | 
else ();  | 
|
| 0 | 211  | 
|
212  | 
fun norm_root ast =  | 
|
213  | 
(case rewrite ast of  | 
|
| 15531 | 214  | 
SOME new_ast => (rewrote ast new_ast; norm_root new_ast)  | 
215  | 
| NONE => ast);  | 
|
| 0 | 216  | 
|
217  | 
fun norm ast =  | 
|
218  | 
(case norm_root ast of  | 
|
219  | 
Appl sub_asts =>  | 
|
220  | 
let  | 
|
221  | 
val old_changes = ! changes;  | 
|
222  | 
val new_ast = Appl (map norm sub_asts);  | 
|
223  | 
in  | 
|
224  | 
if old_changes = ! changes then new_ast else norm_root new_ast  | 
|
225  | 
end  | 
|
226  | 
| atomic_ast => atomic_ast);  | 
|
227  | 
||
228  | 
fun normal ast =  | 
|
229  | 
let  | 
|
230  | 
val old_changes = ! changes;  | 
|
231  | 
val new_ast = norm ast;  | 
|
232  | 
in  | 
|
233  | 
inc passes;  | 
|
234  | 
if old_changes = ! changes then new_ast else normal new_ast  | 
|
235  | 
end;  | 
|
236  | 
||
237  | 
||
| 21962 | 238  | 
    val _ = if trace then tracing ("pre: " ^ str_of_ast pre_ast) else ();
 | 
| 
9372
 
7834e56e2277
AST translation rules no longer require constant head on LHS;
 
wenzelm 
parents: 
8997 
diff
changeset
 | 
239  | 
val post_ast = normal pre_ast;  | 
| 21962 | 240  | 
val _ =  | 
241  | 
if trace orelse stat then  | 
|
242  | 
        tracing ("post: " ^ str_of_ast post_ast ^ "\nnormalize: " ^
 | 
|
243  | 
string_of_int (! passes) ^ " passes, " ^  | 
|
244  | 
string_of_int (! changes) ^ " changes, " ^  | 
|
245  | 
string_of_int (! failed_matches) ^ " matches failed")  | 
|
246  | 
else ();  | 
|
247  | 
in post_ast end;  | 
|
| 0 | 248  | 
|
249  | 
||
| 18 | 250  | 
(* normalize_ast *)  | 
251  | 
||
| 12262 | 252  | 
val trace_ast = ref false;  | 
253  | 
val stat_ast = ref false;  | 
|
254  | 
||
| 18 | 255  | 
fun normalize_ast get_rules ast =  | 
| 8997 | 256  | 
normalize (! trace_ast) (! stat_ast) get_rules ast;  | 
| 18 | 257  | 
|
| 1506 | 258  | 
end;  |