| author | wenzelm | 
| Sat, 14 Jun 2008 17:26:15 +0200 | |
| changeset 27206 | 9a786d5f8821 | 
| parent 26385 | ae7564661e76 | 
| child 28268 | ac8431ecd57e | 
| permissions | -rw-r--r-- | 
| 24584 | 1 | (* Title: Tools/Compute_Oracle/am_compiler.ML | 
| 23174 | 2 | ID: $Id$ | 
| 3 | Author: Steven Obua | |
| 4 | *) | |
| 5 | ||
| 6 | signature COMPILING_AM = | |
| 7 | sig | |
| 8 | include ABSTRACT_MACHINE | |
| 9 | ||
| 23663 | 10 | val set_compiled_rewriter : (term -> term) -> unit | 
| 23174 | 11 | val list_nth : 'a list * int -> 'a | 
| 12 |   val list_map : ('a -> 'b) -> 'a list -> 'b list
 | |
| 13 | end | |
| 14 | ||
| 15 | structure AM_Compiler : COMPILING_AM = struct | |
| 16 | ||
| 17 | val list_nth = List.nth; | |
| 18 | val list_map = map; | |
| 19 | ||
| 23663 | 20 | open AbstractMachine; | 
| 23174 | 21 | |
| 23663 | 22 | val compiled_rewriter = ref (NONE:(term -> term)Option.option) | 
| 23174 | 23 | |
| 24 | fun set_compiled_rewriter r = (compiled_rewriter := SOME r) | |
| 25 | ||
| 26 | type program = (term -> term) | |
| 27 | ||
| 28 | fun count_patternvars PVar = 1 | |
| 29 | | count_patternvars (PConst (_, ps)) = | |
| 30 | List.foldl (fn (p, count) => (count_patternvars p)+count) 0 ps | |
| 31 | ||
| 32 | fun print_rule (p, t) = | |
| 33 | let | |
| 34 | fun str x = Int.toString x | |
| 35 | fun print_pattern n PVar = (n+1, "x"^(str n)) | |
| 36 | | print_pattern n (PConst (c, [])) = (n, "c"^(str c)) | |
| 37 | | print_pattern n (PConst (c, args)) = | |
| 38 | let | |
| 39 | val h = print_pattern n (PConst (c,[])) | |
| 40 | in | |
| 41 | print_pattern_list h args | |
| 42 | end | |
| 43 | and print_pattern_list r [] = r | |
| 44 | | print_pattern_list (n, p) (t::ts) = | |
| 45 | let | |
| 46 | val (n, t) = print_pattern n t | |
| 47 | in | |
| 48 | 		print_pattern_list (n, "App ("^p^", "^t^")") ts
 | |
| 49 | end | |
| 50 | ||
| 51 | val (n, pattern) = print_pattern 0 p | |
| 52 | val pattern = | |
| 53 |             if exists_string Symbol.is_ascii_blank pattern then "(" ^ pattern ^")"
 | |
| 54 | else pattern | |
| 55 | ||
| 56 | fun print_term d (Var x) = (*if x < d then "Var "^(str x) else "x"^(str (n-(x-d)-1))*) | |
| 57 | "Var " ^ str x | |
| 58 | | print_term d (Const c) = "c" ^ str c | |
| 59 | 	  | print_term d (App (a,b)) = "App (" ^ print_term d a ^ ", " ^ print_term d b ^ ")"
 | |
| 60 | 	  | print_term d (Abs c) = "Abs (" ^ print_term (d + 1) c ^ ")"
 | |
| 25217 | 61 | | print_term d (Computed c) = print_term d c | 
| 23174 | 62 | |
| 63 | fun listvars n = if n = 0 then "x0" else "x"^(str n)^", "^(listvars (n-1)) | |
| 64 | ||
| 65 | val term = print_term 0 t | |
| 66 | val term = | |
| 67 | if n > 0 then "Closure (["^(listvars (n-1))^"], "^term^")" | |
| 68 | else "Closure ([], "^term^")" | |
| 69 | ||
| 70 | in | |
| 23663 | 71 | " | weak_reduce (false, stack, "^pattern^") = Continue (false, stack, "^term^")" | 
| 23174 | 72 | end | 
| 73 | ||
| 74 | fun constants_of PVar = [] | |
| 75 | | constants_of (PConst (c, ps)) = c :: maps constants_of ps | |
| 76 | ||
| 77 | fun constants_of_term (Var _) = [] | |
| 78 | | constants_of_term (Abs m) = constants_of_term m | |
| 79 | | constants_of_term (App (a,b)) = (constants_of_term a)@(constants_of_term b) | |
| 80 | | constants_of_term (Const c) = [c] | |
| 25217 | 81 | | constants_of_term (Computed c) = constants_of_term c | 
| 23174 | 82 | |
| 83 | fun load_rules sname name prog = | |
| 84 | let | |
| 85 | val buffer = ref "" | |
| 86 | fun write s = (buffer := (!buffer)^s) | |
| 87 | fun writeln s = (write s; write "\n") | |
| 88 | fun writelist [] = () | |
| 89 | | writelist (s::ss) = (writeln s; writelist ss) | |
| 90 | fun str i = Int.toString i | |
| 91 | val _ = writelist [ | |
| 92 | "structure "^name^" = struct", | |
| 93 | "", | |
| 23663 | 94 | "datatype term = Dummy | App of term * term | Abs of term | Var of int | Const of int | Closure of term list * term"] | 
| 23174 | 95 | val constants = distinct (op =) (maps (fn (p, r) => ((constants_of p)@(constants_of_term r))) prog) | 
| 96 | 	val _ = map (fn x => write (" | c"^(str x))) constants
 | |
| 97 | val _ = writelist [ | |
| 98 | "", | |
| 99 | "datatype stack = SEmpty | SAppL of term * stack | SAppR of term * stack | SAbs of stack", | |
| 23663 | 100 | "", | 
| 101 | "type state = bool * stack * term", | |
| 102 | "", | |
| 103 | "datatype loopstate = Continue of state | Stop of stack * term", | |
| 104 | "", | |
| 105 | "fun proj_C (Continue s) = s", | |
| 106 | " | proj_C _ = raise Match", | |
| 107 | "", | |
| 108 | "fun proj_S (Stop s) = s", | |
| 109 | " | proj_S _ = raise Match", | |
| 110 | "", | |
| 111 | "fun cont (Continue _) = true", | |
| 112 | " | cont _ = false", | |
| 23174 | 113 | "", | 
| 23663 | 114 | "fun do_reduction reduce p =", | 
| 115 | " let", | |
| 116 | " val s = ref (Continue p)", | |
| 117 | " val _ = while cont (!s) do (s := reduce (proj_C (!s)))", | |
| 118 | " in", | |
| 119 | " proj_S (!s)", | |
| 120 | " end", | |
| 121 | ""] | |
| 122 | ||
| 123 | val _ = writelist [ | |
| 124 | "fun weak_reduce (false, stack, Closure (e, App (a, b))) = Continue (false, SAppL (Closure (e, b), stack), Closure (e, a))", | |
| 125 | " | weak_reduce (false, SAppL (b, stack), Closure (e, Abs m)) = Continue (false, stack, Closure (b::e, m))", | |
| 126 | " | weak_reduce (false, stack, c as Closure (e, Abs m)) = Continue (true, stack, c)", | |
| 127 | " | weak_reduce (false, stack, Closure (e, Var n)) = Continue (false, stack, case "^sname^".list_nth (e, n) of Dummy => Var n | r => r)", | |
| 128 | " | weak_reduce (false, stack, Closure (e, c)) = Continue (false, stack, c)"] | |
| 129 | val _ = writelist (map print_rule prog) | |
| 130 | val _ = writelist [ | |
| 131 | " | weak_reduce (false, stack, clos) = Continue (true, stack, clos)", | |
| 132 | " | weak_reduce (true, SAppR (a, stack), b) = Continue (false, stack, App (a,b))", | |
| 133 | " | weak_reduce (true, s as (SAppL (b, stack)), a) = Continue (false, SAppR (a, stack), b)", | |
| 134 | " | weak_reduce (true, stack, c) = Stop (stack, c)", | |
| 23174 | 135 | "", | 
| 23663 | 136 | "fun strong_reduce (false, stack, Closure (e, Abs m)) =", | 
| 23174 | 137 | " let", | 
| 23663 | 138 | " val (stack', wnf) = do_reduction weak_reduce (false, SEmpty, Closure (Dummy::e, m))", | 
| 23174 | 139 | " in", | 
| 23663 | 140 | " case stack' of", | 
| 141 | " SEmpty => Continue (false, SAbs stack, wnf)", | |
| 142 | 		"          | _ => raise ("^sname^".Run \"internal error in strong: weak failed\")",
 | |
| 143 | " end", | |
| 144 | " | strong_reduce (false, stack, clos as (App (u, v))) = Continue (false, SAppL (v, stack), u)", | |
| 145 | " | strong_reduce (false, stack, clos) = Continue (true, stack, clos)", | |
| 146 | " | strong_reduce (true, SAbs stack, m) = Continue (false, stack, Abs m)", | |
| 147 | " | strong_reduce (true, SAppL (b, stack), a) = Continue (false, SAppR (a, stack), b)", | |
| 148 | " | strong_reduce (true, SAppR (a, stack), b) = Continue (true, stack, App (a, b))", | |
| 149 | " | strong_reduce (true, stack, clos) = Stop (stack, clos)", | |
| 23174 | 150 | ""] | 
| 151 | ||
| 152 | val ic = "(case c of "^(implode (map (fn c => (str c)^" => c"^(str c)^" | ") constants))^" _ => Const c)" | |
| 153 | val _ = writelist [ | |
| 154 | 		"fun importTerm ("^sname^".Var x) = Var x",
 | |
| 155 | 		"  | importTerm ("^sname^".Const c) =  "^ic,
 | |
| 156 | 		"  | importTerm ("^sname^".App (a, b)) = App (importTerm a, importTerm b)",
 | |
| 157 | 		"  | importTerm ("^sname^".Abs m) = Abs (importTerm m)",
 | |
| 158 | ""] | |
| 159 | ||
| 23663 | 160 | fun ec c = " | exportTerm c"^(str c)^" = "^sname^".Const "^(str c) | 
| 23174 | 161 | val _ = writelist [ | 
| 23663 | 162 | "fun exportTerm (Var x) = "^sname^".Var x", | 
| 163 | " | exportTerm (Const c) = "^sname^".Const c", | |
| 164 | " | exportTerm (App (a,b)) = "^sname^".App (exportTerm a, exportTerm b)", | |
| 165 | " | exportTerm (Abs m) = "^sname^".Abs (exportTerm m)", | |
| 166 | 		"  | exportTerm (Closure (closlist, clos)) = raise ("^sname^".Run \"internal error, cannot export Closure\")",
 | |
| 167 | 		"  | exportTerm Dummy = raise ("^sname^".Run \"internal error, cannot export Dummy\")"]
 | |
| 23174 | 168 | val _ = writelist (map ec constants) | 
| 169 | ||
| 170 | val _ = writelist [ | |
| 171 | "", | |
| 172 | "fun rewrite t = ", | |
| 173 | " let", | |
| 23663 | 174 | " val (stack, wnf) = do_reduction weak_reduce (false, SEmpty, Closure ([], importTerm t))", | 
| 23174 | 175 | " in", | 
| 176 | " case stack of ", | |
| 23663 | 177 | " SEmpty => (case do_reduction strong_reduce (false, SEmpty, wnf) of", | 
| 23174 | 178 | " (SEmpty, snf) => exportTerm snf", | 
| 179 | 		"                        | _ => raise ("^sname^".Run \"internal error in rewrite: strong failed\"))",
 | |
| 180 | 		"         | _ => (raise ("^sname^".Run \"internal error in rewrite: weak failed\"))",
 | |
| 181 | " end", | |
| 182 | "", | |
| 183 | "val _ = "^sname^".set_compiled_rewriter rewrite", | |
| 184 | "", | |
| 185 | "end;"] | |
| 186 | ||
| 187 | in | |
| 188 | compiled_rewriter := NONE; | |
| 26385 
ae7564661e76
ML runtime compilation: pass position, tuned signature;
 wenzelm parents: 
25520diff
changeset | 189 | use_text (1, "") Output.ml_output false (!buffer); | 
| 23174 | 190 | case !compiled_rewriter of | 
| 191 | NONE => raise (Compile "cannot communicate with compiled function") | |
| 23663 | 192 | | SOME r => (compiled_rewriter := NONE; r) | 
| 23174 | 193 | end | 
| 194 | ||
| 25520 | 195 | fun compile cache_patterns const_arity eqs = | 
| 23174 | 196 | let | 
| 23663 | 197 | 	val _ = if exists (fn (a,b,c) => not (null a)) eqs then raise Compile ("cannot deal with guards") else ()
 | 
| 198 | val eqs = map (fn (a,b,c) => (b,c)) eqs | |
| 199 | 	fun check (p, r) = if check_freevars (count_patternvars p) r then () else raise Compile ("unbound variables in rule") 
 | |
| 23174 | 200 | val _ = map (fn (p, r) => | 
| 23663 | 201 | (check (p, r); | 
| 202 | case p of PVar => raise (Compile "pattern is just a variable") | _ => ())) eqs | |
| 23174 | 203 | in | 
| 204 | load_rules "AM_Compiler" "AM_compiled_code" eqs | |
| 205 | end | |
| 206 | ||
| 207 | fun run prog t = (prog t) | |
| 23663 | 208 | |
| 209 | fun discard p = () | |
| 23174 | 210 | |
| 211 | end | |
| 212 |