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