| author | berghofe | 
| Tue, 27 Sep 2005 12:13:17 +0200 | |
| changeset 17666 | 4708ab4626a5 | 
| parent 16842 | 5979c46853d1 | 
| child 17795 | 5b18c3343028 | 
| permissions | -rw-r--r-- | 
| 16781 | 1 | (* Title: Pure/Tools/am_compiler.ML | 
| 2 | ID: $Id$ | |
| 3 | Author: Steven Obua | |
| 4 | *) | |
| 5 | ||
| 6 | signature COMPILING_AM = | |
| 7 | sig | |
| 16842 | 8 | include ABSTRACT_MACHINE | 
| 16781 | 9 | |
| 16842 | 10 | datatype closure = CVar of int | CConst of int | 
| 11 | | CApp of closure * closure | CAbs of closure | Closure of (closure list) * closure | |
| 16781 | 12 | |
| 16842 | 13 | val set_compiled_rewriter : (term -> closure) -> unit | 
| 14 | val list_nth : 'a list * int -> 'a | |
| 15 |   val list_map : ('a -> 'b) -> 'a list -> 'b list
 | |
| 16781 | 16 | end | 
| 17 | ||
| 18 | structure AM_Compiler :> COMPILING_AM = struct | |
| 19 | ||
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: 
16781diff
changeset | 20 | val list_nth = List.nth; | 
| 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: 
16781diff
changeset | 21 | val list_map = map; | 
| 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: 
16781diff
changeset | 22 | |
| 16781 | 23 | datatype term = Var of int | Const of int | App of term * term | Abs of term | 
| 24 | ||
| 25 | datatype pattern = PVar | PConst of int * (pattern list) | |
| 26 | ||
| 27 | datatype closure = CVar of int | CConst of int | |
| 16842 | 28 | | CApp of closure * closure | CAbs of closure | 
| 29 | | Closure of (closure list) * closure | |
| 16781 | 30 | |
| 31 | val compiled_rewriter = ref (NONE:(term -> closure)Option.option) | |
| 32 | ||
| 33 | fun set_compiled_rewriter r = (compiled_rewriter := SOME r) | |
| 34 | ||
| 35 | type program = (term -> term) | |
| 36 | ||
| 37 | datatype stack = SEmpty | SAppL of closure * stack | SAppR of closure * stack | SAbs of stack | |
| 38 | ||
| 39 | exception Compile of string; | |
| 40 | exception Run of string; | |
| 41 | ||
| 42 | fun clos_of_term (Var x) = CVar x | |
| 43 | | clos_of_term (Const c) = CConst c | |
| 44 | | clos_of_term (App (u, v)) = CApp (clos_of_term u, clos_of_term v) | |
| 45 | | clos_of_term (Abs u) = CAbs (clos_of_term u) | |
| 46 | ||
| 47 | fun term_of_clos (CVar x) = Var x | |
| 48 | | term_of_clos (CConst c) = Const c | |
| 49 | | term_of_clos (CApp (u, v)) = App (term_of_clos u, term_of_clos v) | |
| 50 | | term_of_clos (CAbs u) = Abs (term_of_clos u) | |
| 16842 | 51 | | term_of_clos (Closure (e, u)) = | 
| 52 | raise (Run "internal error: closure in normalized term found") | |
| 16781 | 53 | |
| 54 | fun strip_closure args (CApp (a,b)) = strip_closure (b::args) a | |
| 55 | | strip_closure args x = (x, args) | |
| 56 | ||
| 16842 | 57 | (*Returns true iff at most 0 .. (free-1) occur unbound. therefore | 
| 58 | check_freevars 0 t iff t is closed*) | |
| 16781 | 59 | fun check_freevars free (Var x) = x < free | 
| 60 | | check_freevars free (Const c) = true | |
| 61 | | check_freevars free (App (u, v)) = check_freevars free u andalso check_freevars free v | |
| 62 | | check_freevars free (Abs m) = check_freevars (free+1) m | |
| 63 | ||
| 64 | fun count_patternvars PVar = 1 | |
| 16842 | 65 | | count_patternvars (PConst (_, ps)) = | 
| 66 | List.foldl (fn (p, count) => (count_patternvars p)+count) 0 ps | |
| 16781 | 67 | |
| 68 | fun print_rule (p, t) = | |
| 69 | let | |
| 70 | fun str x = Int.toString x | |
| 71 | fun print_pattern n PVar = (n+1, "x"^(str n)) | |
| 72 | | print_pattern n (PConst (c, [])) = (n, "c"^(str c)) | |
| 73 | | print_pattern n (PConst (c, args)) = | |
| 74 | let | |
| 75 | val h = print_pattern n (PConst (c,[])) | |
| 76 | in | |
| 77 | print_pattern_list h args | |
| 78 | end | |
| 79 | and print_pattern_list r [] = r | |
| 80 | | print_pattern_list (n, p) (t::ts) = | |
| 81 | let | |
| 82 | val (n, t) = print_pattern n t | |
| 83 | in | |
| 84 | 		print_pattern_list (n, "App ("^p^", "^t^")") ts
 | |
| 85 | end | |
| 86 | ||
| 87 | val (n, pattern) = print_pattern 0 p | |
| 16842 | 88 | val pattern = | 
| 89 |             if List.exists Char.isSpace (String.explode pattern) then "("^pattern^")"
 | |
| 90 | else pattern | |
| 16781 | 91 | |
| 16842 | 92 | fun print_term d (Var x) = (*if x < d then "Var "^(str x) else "x"^(str (n-(x-d)-1))*) | 
| 93 | "Var " ^ str x | |
| 94 | | print_term d (Const c) = "c" ^ str c | |
| 95 | 	  | print_term d (App (a,b)) = "App (" ^ print_term d a ^ ", " ^ print_term d b ^ ")"
 | |
| 96 | 	  | print_term d (Abs c) = "Abs (" ^ print_term (d + 1) c ^ ")"
 | |
| 16781 | 97 | |
| 98 | fun listvars n = if n = 0 then "x0" else "x"^(str n)^", "^(listvars (n-1)) | |
| 99 | ||
| 100 | val term = print_term 0 t | |
| 16842 | 101 | val term = | 
| 102 | if n > 0 then "Closure (["^(listvars (n-1))^"], "^term^")" | |
| 103 | else "Closure ([], "^term^")" | |
| 16781 | 104 | |
| 105 | in | |
| 106 | 	"lookup stack "^pattern^" = weak stack ("^term^")"
 | |
| 107 | end | |
| 108 | ||
| 109 | fun remove_dups (c::cs) = | |
| 110 | let val cs = remove_dups cs in | |
| 111 | if (List.exists (fn x => x=c) cs) then cs else c::cs | |
| 112 | end | |
| 113 | | remove_dups [] = [] | |
| 114 | ||
| 115 | fun constants_of PVar = [] | |
| 16842 | 116 | | constants_of (PConst (c, ps)) = c :: List.concat (map constants_of ps) | 
| 16781 | 117 | |
| 118 | fun constants_of_term (Var _) = [] | |
| 119 | | constants_of_term (Abs m) = constants_of_term m | |
| 120 | | constants_of_term (App (a,b)) = (constants_of_term a)@(constants_of_term b) | |
| 121 | | constants_of_term (Const c) = [c] | |
| 122 | ||
| 123 | fun load_rules sname name prog = | |
| 124 | let | |
| 16842 | 125 | (* FIXME consider using more readable/efficient Buffer.empty |> fold Buffer.add etc. *) | 
| 16781 | 126 | val buffer = ref "" | 
| 127 | fun write s = (buffer := (!buffer)^s) | |
| 128 | fun writeln s = (write s; write "\n") | |
| 129 | fun writelist [] = () | |
| 130 | | writelist (s::ss) = (writeln s; writelist ss) | |
| 131 | fun str i = Int.toString i | |
| 132 | val _ = writelist [ | |
| 133 | "structure "^name^" = struct", | |
| 134 | "", | |
| 135 | "datatype term = App of term * term | Abs of term | Var of int | Const of int | Closure of term list * term"] | |
| 136 | val constants = remove_dups (List.concat (map (fn (p, r) => ((constants_of p)@(constants_of_term r))) prog)) | |
| 137 | 	val _ = map (fn x => write (" | c"^(str x))) constants
 | |
| 138 | val _ = writelist [ | |
| 139 | "", | |
| 140 | "datatype stack = SEmpty | SAppL of term * stack | SAppR of term * stack | SAbs of stack", | |
| 141 | ""] | |
| 142 | val _ = (case prog of | |
| 143 | 		    r::rs => (writeln ("fun "^(print_rule r)); 
 | |
| 144 | 			      map (fn r => writeln("  | "^(print_rule r))) rs; 
 | |
| 145 | 			      writeln ("  | lookup stack clos = weak_last stack clos"); ())								
 | |
| 146 | | [] => (writeln "fun lookup stack clos = weak_last stack clos")) | |
| 147 | val _ = writelist [ | |
| 148 | "and weak stack (Closure (e, App (a, b))) = weak (SAppL (Closure (e, b), stack)) (Closure (e, a))", | |
| 149 | " | weak (SAppL (b, stack)) (Closure (e, Abs m)) = weak stack (Closure (b::e, m))", | |
| 150 | " | weak stack (clos as Closure (_, Abs _)) = weak_last stack clos", | |
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: 
16781diff
changeset | 151 | 		"  | weak stack (Closure (e, Var n)) = weak stack ("^sname^".list_nth (e, n) handle _ => (Var (n-(length e))))",
 | 
| 16781 | 152 | " | weak stack (Closure (e, c)) = weak stack c", | 
| 153 | " | weak stack clos = lookup stack clos", | |
| 154 | "and weak_last (SAppR (a, stack)) b = weak stack (App(a, b))", | |
| 155 | " | weak_last (SAppL (b, stack)) a = weak (SAppR (a, stack)) b", | |
| 156 | " | weak_last stack c = (stack, c)", | |
| 157 | "", | |
| 158 | "fun lift n (v as Var m) = if m < n then v else Var (m+1)", | |
| 159 | " | lift n (Abs t) = Abs (lift (n+1) t)", | |
| 160 | " | lift n (App (a,b)) = App (lift n a, lift n b)", | |
| 161 | " | lift n (Closure (e, a)) = Closure (lift_env n e, lift (n+(length e)) a)", | |
| 162 | " | lift n c = c", | |
| 163 | "and lift_env n e = map (lift n) e", | |
| 164 | "", | |
| 165 | "fun strong stack (Closure (e, Abs m)) = ", | |
| 166 | " let", | |
| 167 | " val (stack', wnf) = weak SEmpty (Closure ((Var 0)::(lift_env 0 e), m))", | |
| 168 | " in", | |
| 169 | " case stack' of", | |
| 170 | " SEmpty => strong (SAbs stack) wnf", | |
| 171 | 		"         | _ => raise ("^sname^".Run \"internal error in strong: weak failed\")",
 | |
| 172 | " end", | |
| 173 | " | strong stack (clos as (App (u, v))) = strong (SAppL (v, stack)) u", | |
| 174 | " | strong stack clos = strong_last stack clos", | |
| 175 | "and strong_last (SAbs stack) m = strong stack (Abs m)", | |
| 176 | " | strong_last (SAppL (b, stack)) a = strong (SAppR (a, stack)) b", | |
| 177 | " | strong_last (SAppR (a, stack)) b = strong_last stack (App (a, b))", | |
| 178 | " | strong_last stack clos = (stack, clos)", | |
| 179 | ""] | |
| 180 | ||
| 181 | val ic = "(case c of "^(String.concat (map (fn c => (str c)^" => c"^(str c)^" | ") constants))^" _ => Const c)" | |
| 182 | val _ = writelist [ | |
| 183 | 		"fun importTerm ("^sname^".Var x) = Var x",
 | |
| 184 | 		"  | importTerm ("^sname^".Const c) =  "^ic,
 | |
| 185 | 		"  | importTerm ("^sname^".App (a, b)) = App (importTerm a, importTerm b)",
 | |
| 186 | 		"  | importTerm ("^sname^".Abs m) = Abs (importTerm m)",
 | |
| 187 | ""] | |
| 188 | ||
| 189 | fun ec c = " | exportTerm c"^(str c)^" = "^sname^".CConst "^(str c) | |
| 190 | val _ = writelist [ | |
| 191 | "fun exportTerm (Var x) = "^sname^".CVar x", | |
| 192 | " | exportTerm (Const c) = "^sname^".CConst c", | |
| 193 | " | exportTerm (App (a,b)) = "^sname^".CApp (exportTerm a, exportTerm b)", | |
| 194 | " | exportTerm (Abs m) = "^sname^".CAbs (exportTerm m)", | |
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: 
16781diff
changeset | 195 | 		"  | exportTerm (Closure (closlist, clos)) = "^sname^".Closure ("^sname^".list_map exportTerm closlist, exportTerm clos)"]
 | 
| 16781 | 196 | val _ = writelist (map ec constants) | 
| 197 | ||
| 198 | val _ = writelist [ | |
| 199 | "", | |
| 200 | "fun rewrite t = ", | |
| 201 | " let", | |
| 202 | " val (stack, wnf) = weak SEmpty (Closure ([], importTerm t))", | |
| 203 | " in", | |
| 204 | " case stack of ", | |
| 205 | " SEmpty => (case strong SEmpty wnf of", | |
| 206 | " (SEmpty, snf) => exportTerm snf", | |
| 207 | 		"                        | _ => raise ("^sname^".Run \"internal error in rewrite: strong failed\"))",
 | |
| 208 | 		"         | _ => (raise ("^sname^".Run \"internal error in rewrite: weak failed\"))",
 | |
| 209 | " end", | |
| 210 | "", | |
| 211 | "val _ = "^sname^".set_compiled_rewriter rewrite", | |
| 212 | "", | |
| 213 | "end;"] | |
| 214 | ||
| 215 | val _ = | |
| 216 | let | |
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: 
16781diff
changeset | 217 | (*val fout = TextIO.openOut "gen_code.ML" | 
| 16781 | 218 | val _ = TextIO.output (fout, !buffer) | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: 
16781diff
changeset | 219 | val _ = TextIO.closeOut fout*) | 
| 16781 | 220 | in | 
| 221 | () | |
| 222 | end | |
| 223 | in | |
| 224 | compiled_rewriter := NONE; | |
| 16842 | 225 | use_text (K (), K ()) false (!buffer); | 
| 16781 | 226 | case !compiled_rewriter of | 
| 227 | NONE => raise (Compile "cannot communicate with compiled function") | |
| 228 | | SOME r => (compiled_rewriter := NONE; fn t => term_of_clos (r t)) | |
| 229 | end | |
| 230 | ||
| 231 | fun compile eqs = | |
| 232 | let | |
| 233 | val _ = map (fn (p, r) => | |
| 234 | (check_freevars (count_patternvars p) r; | |
| 235 | case p of PVar => raise (Compile "pattern reduces to a variable") | _ => ())) eqs | |
| 236 | in | |
| 237 | load_rules "AM_Compiler" "AM_compiled_code" eqs | |
| 238 | end | |
| 239 | ||
| 240 | fun run prog t = (prog t) | |
| 241 | ||
| 242 | end | |
| 243 | ||
| 244 | structure AbstractMachine = AM_Compiler |