obua@16781: (* Title: Pure/Tools/am_interpreter.ML obua@16781: ID: $Id$ obua@16781: Author: Steven Obua obua@16781: *) obua@16781: wenzelm@16842: signature ABSTRACT_MACHINE = wenzelm@16842: sig wenzelm@16842: obua@16781: datatype term = Var of int | Const of int | App of term * term | Abs of term obua@16781: obua@16781: datatype pattern = PVar | PConst of int * (pattern list) obua@16781: obua@16781: type program obua@16781: obua@16781: exception Compile of string; obua@16781: val compile : (pattern * term) list -> program wenzelm@16842: obua@16781: exception Run of string; obua@16781: val run : program -> term -> term obua@16781: obua@16781: end obua@16781: obua@16781: structure AM_Interpreter :> ABSTRACT_MACHINE = struct obua@16781: obua@16781: datatype term = Var of int | Const of int | App of term * term | Abs of term obua@16781: obua@16781: datatype pattern = PVar | PConst of int * (pattern list) obua@16781: obua@16781: datatype closure = CVar of int | CConst of int wenzelm@16842: | CApp of closure * closure | CAbs of closure wenzelm@16842: | Closure of (closure list) * closure obua@16781: obua@16782: structure prog_struct = TableFun(type key = int*int val ord = prod_ord int_ord int_ord); obua@16781: obua@16782: type program = ((pattern * closure) list) prog_struct.table obua@16781: obua@16781: datatype stack = SEmpty | SAppL of closure * stack | SAppR of closure * stack | SAbs of stack obua@16781: obua@16781: exception Compile of string; obua@16781: exception Run of string; obua@16781: obua@16781: fun clos_of_term (Var x) = CVar x obua@16781: | clos_of_term (Const c) = CConst c obua@16781: | clos_of_term (App (u, v)) = CApp (clos_of_term u, clos_of_term v) obua@16781: | clos_of_term (Abs u) = CAbs (clos_of_term u) obua@16781: obua@16781: fun term_of_clos (CVar x) = Var x obua@16781: | term_of_clos (CConst c) = Const c obua@16781: | term_of_clos (CApp (u, v)) = App (term_of_clos u, term_of_clos v) obua@16781: | term_of_clos (CAbs u) = Abs (term_of_clos u) obua@16781: | term_of_clos (Closure (e, u)) = raise (Run "internal error: closure in normalized term found") obua@16781: obua@16781: fun strip_closure args (CApp (a,b)) = strip_closure (b::args) a obua@16781: | strip_closure args x = (x, args) obua@16781: obua@16781: fun len_head_of_closure n (CApp (a,b)) = len_head_of_closure (n+1) a obua@16781: | len_head_of_closure n x = (n, x) obua@16781: obua@16781: obua@16781: (* earlier occurrence of PVar corresponds to higher de Bruijn index *) obua@16781: fun pattern_match args PVar clos = SOME (clos::args) wenzelm@16842: | pattern_match args (PConst (c, patterns)) clos = obua@16781: let wenzelm@16842: val (f, closargs) = strip_closure [] clos obua@16781: in wenzelm@16842: case f of wenzelm@16842: CConst d => wenzelm@16842: if c = d then wenzelm@16842: pattern_match_list args patterns closargs wenzelm@16842: else wenzelm@16842: NONE wenzelm@16842: | _ => NONE obua@16781: end obua@16781: and pattern_match_list args [] [] = SOME args wenzelm@16842: | pattern_match_list args (p::ps) (c::cs) = obua@16781: (case pattern_match args p c of wenzelm@16842: NONE => NONE obua@16781: | SOME args => pattern_match_list args ps cs) obua@16781: | pattern_match_list _ _ _ = NONE obua@16781: obua@16781: (* Returns true iff at most 0 .. (free-1) occur unbound. therefore check_freevars 0 t iff t is closed *) obua@16781: fun check_freevars free (Var x) = x < free obua@16781: | check_freevars free (Const c) = true obua@16781: | check_freevars free (App (u, v)) = check_freevars free u andalso check_freevars free v obua@16781: | check_freevars free (Abs m) = check_freevars (free+1) m obua@16781: obua@16781: fun count_patternvars PVar = 1 obua@16781: | count_patternvars (PConst (_, ps)) = List.foldl (fn (p, count) => (count_patternvars p)+count) 0 ps obua@16781: obua@16781: fun pattern_key (PConst (c, ps)) = (c, length ps) obua@16781: | pattern_key _ = raise (Compile "pattern reduces to variable") obua@16781: wenzelm@16842: fun compile eqs = obua@16781: let wenzelm@16842: val eqs = map (fn (p, r) => (check_freevars (count_patternvars p) r; wenzelm@16842: (pattern_key p, (p, clos_of_term r)))) eqs obua@16781: in wenzelm@16842: prog_struct.make (map (fn (k, a) => (k, [a])) eqs) wenzelm@16842: end obua@16781: obua@16781: fun match_rules n [] clos = NONE obua@16781: | match_rules n ((p,eq)::rs) clos = obua@16781: case pattern_match [] p clos of wenzelm@16842: NONE => match_rules (n+1) rs clos obua@16781: | SOME args => SOME (Closure (args, eq)) obua@16781: wenzelm@16842: fun match_closure prog clos = obua@16781: case len_head_of_closure 0 clos of wenzelm@16842: (len, CConst c) => wenzelm@16842: (case prog_struct.lookup (prog, (c, len)) of wenzelm@16842: NONE => NONE wenzelm@16842: | SOME rules => match_rules 0 rules clos) obua@16781: | _ => NONE obua@16781: obua@16781: fun lift n (c as (CConst _)) = c obua@16781: | lift n (v as CVar m) = if m < n then v else CVar (m+1) obua@16781: | lift n (CAbs t) = CAbs (lift (n+1) t) obua@16781: | lift n (CApp (a,b)) = CApp (lift n a, lift n b) obua@16781: | lift n (Closure (e, a)) = Closure (lift_env n e, lift (n+(length e)) a) obua@16781: and lift_env n e = map (lift n) e obua@16781: obua@16781: fun weak prog stack (Closure (e, CApp (a, b))) = weak prog (SAppL (Closure (e, b), stack)) (Closure (e, a)) obua@16781: | weak prog (SAppL (b, stack)) (Closure (e, CAbs m)) = weak prog stack (Closure (b::e, m)) obua@16781: | weak prog stack (Closure (e, CVar n)) = weak prog stack (List.nth (e, n) handle Subscript => (CVar (n-(length e)))) obua@16781: | weak prog stack (Closure (e, c as CConst _)) = weak prog stack c wenzelm@16842: | weak prog stack clos = obua@16781: case match_closure prog clos of wenzelm@16842: NONE => weak_last prog stack clos obua@16781: | SOME r => weak prog stack r obua@16781: and weak_last prog (SAppR (a, stack)) b = weak prog stack (CApp (a,b)) obua@16781: | weak_last prog (s as (SAppL (b, stack))) a = weak prog (SAppR (a, stack)) b wenzelm@16842: | weak_last prog stack c = (stack, c) obua@16781: wenzelm@16842: fun strong prog stack (Closure (e, CAbs m)) = obua@16781: let wenzelm@16842: val (stack', wnf) = weak prog SEmpty (Closure ((CVar 0)::(lift_env 0 e), m)) obua@16781: in wenzelm@16842: case stack' of wenzelm@16842: SEmpty => strong prog (SAbs stack) wnf wenzelm@16842: | _ => raise (Run "internal error in strong: weak failed") obua@16781: end obua@16781: | strong prog stack (clos as (CApp (u, v))) = strong prog (SAppL (v, stack)) u obua@16781: | strong prog stack clos = strong_last prog stack clos obua@16781: and strong_last prog (SAbs stack) m = strong prog stack (CAbs m) obua@16781: | strong_last prog (SAppL (b, stack)) a = strong prog (SAppR (a, stack)) b obua@16781: | strong_last prog (SAppR (a, stack)) b = strong_last prog stack (CApp (a, b)) obua@16781: | strong_last prog stack clos = (stack, clos) obua@16781: wenzelm@16842: fun run prog t = obua@16781: let wenzelm@16842: val (stack, wnf) = weak prog SEmpty (Closure ([], clos_of_term t)) obua@16781: in wenzelm@16842: case stack of wenzelm@16842: SEmpty => (case strong prog SEmpty wnf of wenzelm@16842: (SEmpty, snf) => term_of_clos snf wenzelm@16842: | _ => raise (Run "internal error in run: strong failed")) wenzelm@16842: | _ => raise (Run "internal error in run: weak failed") obua@16781: end wenzelm@16842: obua@16781: end obua@16781: obua@16781: structure AbstractMachine = AM_Interpreter