haftmann@24590: (* Title: Tools/nbe.ML haftmann@24155: ID: $Id$ haftmann@24155: Authors: Klaus Aehlig, LMU Muenchen; Tobias Nipkow, Florian Haftmann, TU Muenchen haftmann@24155: haftmann@24839: Normalization by evaluation, based on generic code generator. haftmann@24155: *) haftmann@24155: haftmann@24155: signature NBE = haftmann@24155: sig haftmann@25101: val norm_conv: cterm -> thm haftmann@25101: val norm_term: theory -> term -> term haftmann@25101: wenzelm@25204: datatype Univ = haftmann@26064: Const of int * Univ list (*named (uninterpreted) constants*) haftmann@25924: | Free of string * Univ list (*free (uninterpreted) variables*) haftmann@25924: | DFree of string * int (*free (uninterpreted) dictionary parameters*) haftmann@24155: | BVar of int * Univ list haftmann@24155: | Abs of (int * (Univ list -> Univ)) * Univ list; haftmann@25924: val apps: Univ -> Univ list -> Univ (*explicit applications*) haftmann@25944: val abss: int -> (Univ list -> Univ) -> Univ haftmann@25924: (*abstractions as closures*) haftmann@24155: wenzelm@25204: val univs_ref: (unit -> Univ list -> Univ list) option ref haftmann@24155: val trace: bool ref haftmann@25924: haftmann@26747: val setup: theory -> theory haftmann@24155: end; haftmann@24155: haftmann@24155: structure Nbe: NBE = haftmann@24155: struct haftmann@24155: haftmann@24155: (* generic non-sense *) haftmann@24155: haftmann@24155: val trace = ref false; haftmann@24155: fun tracing f x = if !trace then (Output.tracing (f x); x) else x; haftmann@24155: haftmann@24155: haftmann@24155: (** the semantical universe **) haftmann@24155: haftmann@24155: (* haftmann@24155: Functions are given by their semantical function value. To avoid haftmann@24155: trouble with the ML-type system, these functions have the most haftmann@24155: generic type, that is "Univ list -> Univ". The calling convention is haftmann@24155: that the arguments come as a list, the last argument first. In haftmann@24155: other words, a function call that usually would look like haftmann@24155: haftmann@24155: f x_1 x_2 ... x_n or f(x_1,x_2, ..., x_n) haftmann@24155: haftmann@24155: would be in our convention called as haftmann@24155: haftmann@24155: f [x_n,..,x_2,x_1] haftmann@24155: haftmann@24155: Moreover, to handle functions that are still waiting for some haftmann@24155: arguments we have additionally a list of arguments collected to far haftmann@24155: and the number of arguments we're still waiting for. haftmann@24155: *) haftmann@24155: wenzelm@25204: datatype Univ = haftmann@26064: Const of int * Univ list (*named (uninterpreted) constants*) haftmann@24155: | Free of string * Univ list (*free variables*) haftmann@25924: | DFree of string * int (*free (uninterpreted) dictionary parameters*) haftmann@27499: | BVar of int * Univ list (*bound variables, named*) haftmann@24155: | Abs of (int * (Univ list -> Univ)) * Univ list haftmann@27499: (*abstractions as closures*); haftmann@24155: haftmann@24155: (* constructor functions *) haftmann@24155: haftmann@25944: fun abss n f = Abs ((n, f), []); haftmann@25924: fun apps (Abs ((n, f), xs)) ys = let val k = n - length ys in haftmann@27499: case int_ord (k, 0) haftmann@27499: of EQUAL => f (ys @ xs) haftmann@27499: | LESS => let val (zs, ws) = chop (~ k) ys in apps (f (ws @ xs)) zs end haftmann@27499: | GREATER => Abs ((k, f), ys @ xs) (*note: reverse convention also for apps!*) haftmann@27499: end haftmann@25924: | apps (Const (name, xs)) ys = Const (name, ys @ xs) haftmann@25924: | apps (Free (name, xs)) ys = Free (name, ys @ xs) haftmann@27499: | apps (BVar (n, xs)) ys = BVar (n, ys @ xs); haftmann@24155: haftmann@24155: haftmann@24155: (** assembling and compiling ML code from terms **) haftmann@24155: haftmann@24155: (* abstract ML syntax *) haftmann@24155: haftmann@24155: infix 9 `$` `$$`; haftmann@24155: fun e1 `$` e2 = "(" ^ e1 ^ " " ^ e2 ^ ")"; haftmann@25101: fun e `$$` [] = e haftmann@25101: | e `$$` es = "(" ^ e ^ " " ^ space_implode " " es ^ ")"; haftmann@24590: fun ml_abs v e = "(fn " ^ v ^ " => " ^ e ^ ")"; haftmann@24155: haftmann@24155: fun ml_cases t cs = haftmann@24155: "(case " ^ t ^ " of " ^ space_implode " | " (map (fn (p, t) => p ^ " => " ^ t) cs) ^ ")"; haftmann@25944: fun ml_Let d e = "let\n" ^ d ^ " in " ^ e ^ " end"; haftmann@24155: haftmann@24155: fun ml_list es = "[" ^ commas es ^ "]"; haftmann@24155: haftmann@24155: fun ml_fundefs ([(name, [([], e)])]) = haftmann@24155: "val " ^ name ^ " = " ^ e ^ "\n" haftmann@24155: | ml_fundefs (eqs :: eqss) = haftmann@24155: let haftmann@24155: fun fundef (name, eqs) = haftmann@24155: let haftmann@24155: fun eqn (es, e) = name ^ " " ^ space_implode " " es ^ " = " ^ e haftmann@24155: in space_implode "\n | " (map eqn eqs) end; haftmann@24155: in haftmann@24155: (prefix "fun " o fundef) eqs :: map (prefix "and " o fundef) eqss wenzelm@26931: |> cat_lines haftmann@24155: |> suffix "\n" haftmann@24155: end; haftmann@24155: haftmann@25944: (* nbe specific syntax and sandbox communication *) haftmann@25944: haftmann@25944: val univs_ref = ref (NONE : (unit -> Univ list -> Univ list) option); haftmann@24155: haftmann@24155: local haftmann@24155: val prefix = "Nbe."; haftmann@25944: val name_ref = prefix ^ "univs_ref"; haftmann@24155: val name_const = prefix ^ "Const"; haftmann@25944: val name_abss = prefix ^ "abss"; haftmann@25924: val name_apps = prefix ^ "apps"; haftmann@24155: in haftmann@24155: haftmann@25944: val univs_cookie = (name_ref, univs_ref); haftmann@25944: haftmann@25935: fun nbe_fun "" = "nbe_value" haftmann@25935: | nbe_fun c = "c_" ^ translate_string (fn "." => "_" | c => c) c; haftmann@25101: fun nbe_dict v n = "d_" ^ v ^ "_" ^ string_of_int n; haftmann@24155: fun nbe_bound v = "v_" ^ v; haftmann@24155: haftmann@25924: (*note: these three are the "turning spots" where proper argument order is established!*) haftmann@25924: fun nbe_apps t [] = t haftmann@25924: | nbe_apps t ts = name_apps `$$` [t, ml_list (rev ts)]; haftmann@25924: fun nbe_apps_local c ts = nbe_fun c `$` ml_list (rev ts); haftmann@26064: fun nbe_apps_constr idx ts = haftmann@26064: name_const `$` ("(" ^ string_of_int idx ^ ", " ^ ml_list (rev ts) ^ ")"); haftmann@25924: haftmann@24155: fun nbe_abss 0 f = f `$` ml_list [] haftmann@25944: | nbe_abss n f = name_abss `$$` [string_of_int n, f]; haftmann@24155: haftmann@24155: end; haftmann@24155: haftmann@24219: open BasicCodeThingol; haftmann@24155: haftmann@25865: (* code generation *) haftmann@24155: haftmann@26064: fun assemble_eqnss idx_of deps eqnss = haftmann@25944: let haftmann@25944: fun prep_eqns (c, (vs, eqns)) = haftmann@25924: let haftmann@25944: val dicts = maps (fn (v, sort) => map_index (nbe_dict v o fst) sort) vs; haftmann@25944: val num_args = length dicts + (length o fst o hd) eqns; haftmann@25944: in (c, (num_args, (dicts, eqns))) end; haftmann@25944: val eqnss' = map prep_eqns eqnss; haftmann@25944: haftmann@25944: fun assemble_constapp c dss ts = haftmann@25944: let haftmann@25944: val ts' = (maps o map) assemble_idict dss @ ts; haftmann@25944: in case AList.lookup (op =) eqnss' c haftmann@25944: of SOME (n, _) => if n <= length ts' haftmann@25924: then let val (ts1, ts2) = chop n ts' haftmann@25924: in nbe_apps (nbe_apps_local c ts1) ts2 haftmann@25924: end else nbe_apps (nbe_abss n (nbe_fun c)) ts' haftmann@25944: | NONE => if member (op =) deps c haftmann@25944: then nbe_apps (nbe_fun c) ts' haftmann@26064: else nbe_apps_constr (idx_of c) ts' haftmann@25924: end haftmann@25944: and assemble_idict (DictConst (inst, dss)) = haftmann@25944: assemble_constapp inst dss [] haftmann@25944: | assemble_idict (DictVar (supers, (v, (n, _)))) = haftmann@25944: fold_rev (fn super => assemble_constapp super [] o single) supers (nbe_dict v n); haftmann@25924: haftmann@25944: fun assemble_iterm constapp = haftmann@24155: let haftmann@25944: fun of_iterm t = haftmann@25944: let haftmann@25944: val (t', ts) = CodeThingol.unfold_app t haftmann@25944: in of_iapp t' (fold_rev (cons o of_iterm) ts []) end haftmann@25944: and of_iapp (IConst (c, (dss, _))) ts = constapp c dss ts haftmann@25944: | of_iapp (IVar v) ts = nbe_apps (nbe_bound v) ts haftmann@25944: | of_iapp ((v, _) `|-> t) ts = haftmann@25944: nbe_apps (nbe_abss 1 (ml_abs (ml_list [nbe_bound v]) (of_iterm t))) ts haftmann@25944: | of_iapp (ICase (((t, _), cs), t0)) ts = haftmann@25944: nbe_apps (ml_cases (of_iterm t) (map (pairself of_iterm) cs haftmann@25944: @ [("_", of_iterm t0)])) ts haftmann@25944: in of_iterm end; haftmann@24155: haftmann@25944: fun assemble_eqns (c, (num_args, (dicts, eqns))) = haftmann@25944: let haftmann@26064: val assemble_arg = assemble_iterm haftmann@26064: (fn c => fn _ => fn ts => nbe_apps_constr (idx_of c) ts); haftmann@25944: val assemble_rhs = assemble_iterm assemble_constapp; haftmann@25944: fun assemble_eqn (args, rhs) = haftmann@25944: ([ml_list (rev (dicts @ map assemble_arg args))], assemble_rhs rhs); haftmann@25944: val default_args = map nbe_bound (Name.invent_list [] "a" num_args); haftmann@26064: val default_eqn = if c = "" then NONE haftmann@26064: else SOME ([ml_list (rev default_args)], haftmann@26064: nbe_apps_constr (idx_of c) default_args); haftmann@25944: in haftmann@26064: ((nbe_fun c, map assemble_eqn eqns @ the_list default_eqn), haftmann@25944: nbe_abss num_args (nbe_fun c)) haftmann@25944: end; haftmann@24155: haftmann@25944: val (fun_vars, fun_vals) = map_split assemble_eqns eqnss'; haftmann@25944: val deps_vars = ml_list (map nbe_fun deps); haftmann@25944: in ml_abs deps_vars (ml_Let (ml_fundefs fun_vars) (ml_list fun_vals)) end; haftmann@25944: haftmann@25944: (* code compilation *) haftmann@25944: haftmann@25944: fun compile_eqnss gr raw_deps [] = [] haftmann@25944: | compile_eqnss gr raw_deps eqnss = haftmann@24155: let haftmann@26064: val (deps, deps_vals) = split_list (map_filter haftmann@26064: (fn dep => Option.map (fn univ => (dep, univ)) (fst ((Graph.get_node gr dep)))) raw_deps); haftmann@26064: val idx_of = raw_deps haftmann@26064: |> map (fn dep => (dep, snd (Graph.get_node gr dep))) haftmann@26064: |> AList.lookup (op =) haftmann@26064: |> (fn f => the o f); haftmann@26064: val s = assemble_eqnss idx_of deps eqnss; haftmann@24155: val cs = map fst eqnss; haftmann@25944: in haftmann@25944: s haftmann@25944: |> tracing (fn s => "\n--- code to be evaluated:\n" ^ s) haftmann@25944: |> ML_Context.evaluate haftmann@25944: (Output.tracing o enclose "\n---compiler echo:\n" "\n---\n", haftmann@25944: Output.tracing o enclose "\n--- compiler echo (with error):\n" "\n---\n") haftmann@25944: (!trace) univs_cookie haftmann@25944: |> (fn f => f deps_vals) haftmann@25944: |> (fn univs => cs ~~ univs) haftmann@25944: end; haftmann@25190: haftmann@25944: (* preparing function equations *) haftmann@24155: haftmann@25101: fun eqns_of_stmt (_, CodeThingol.Fun (_, [])) = haftmann@25101: [] haftmann@25101: | eqns_of_stmt (const, CodeThingol.Fun ((vs, _), eqns)) = haftmann@25101: [(const, (vs, map fst eqns))] haftmann@25101: | eqns_of_stmt (_, CodeThingol.Datatypecons _) = haftmann@25101: [] haftmann@25101: | eqns_of_stmt (_, CodeThingol.Datatype _) = haftmann@25101: [] haftmann@25101: | eqns_of_stmt (class, CodeThingol.Class (v, (superclasses, classops))) = haftmann@25101: let haftmann@25101: val names = map snd superclasses @ map fst classops; haftmann@25101: val params = Name.invent_list [] "d" (length names); haftmann@25101: fun mk (k, name) = haftmann@25101: (name, ([(v, [])], haftmann@25101: [([IConst (class, ([], [])) `$$ map IVar params], IVar (nth params k))])); haftmann@25101: in map_index mk names end haftmann@25101: | eqns_of_stmt (_, CodeThingol.Classrel _) = haftmann@25101: [] haftmann@25101: | eqns_of_stmt (_, CodeThingol.Classparam _) = haftmann@25101: [] haftmann@25101: | eqns_of_stmt (inst, CodeThingol.Classinst ((class, (_, arities)), (superinsts, instops))) = haftmann@25101: [(inst, (arities, [([], IConst (class, ([], [])) `$$ haftmann@25101: map (fn (_, (_, (inst, dicts))) => IConst (inst, (dicts, []))) superinsts haftmann@25101: @ map (IConst o snd o fst) instops)]))]; haftmann@24155: haftmann@25101: fun compile_stmts stmts_deps = haftmann@25101: let haftmann@25101: val names = map (fst o fst) stmts_deps; haftmann@25101: val names_deps = map (fn ((name, _), deps) => (name, deps)) stmts_deps; haftmann@25101: val eqnss = maps (eqns_of_stmt o fst) stmts_deps; haftmann@26064: val refl_deps = names_deps haftmann@25190: |> maps snd haftmann@25190: |> distinct (op =) haftmann@26064: |> fold (insert (op =)) names; haftmann@26064: fun new_node name (gr, (maxidx, idx_tab)) = if can (Graph.get_node gr) name haftmann@26064: then (gr, (maxidx, idx_tab)) haftmann@26064: else (Graph.new_node (name, (NONE, maxidx)) gr, haftmann@26064: (maxidx + 1, Inttab.update_new (maxidx, name) idx_tab)); haftmann@25190: fun compile gr = eqnss haftmann@26064: |> compile_eqnss gr refl_deps haftmann@25190: |> rpair gr; haftmann@25101: in haftmann@26064: fold new_node refl_deps haftmann@26064: #> apfst (fold (fn (name, deps) => fold (curry Graph.add_edge name) deps) names_deps haftmann@26064: #> compile haftmann@26064: #-> fold (fn (name, univ) => (Graph.map_node name o apfst) (K (SOME univ)))) haftmann@25101: end; haftmann@24155: haftmann@27103: fun ensure_stmts program = haftmann@25101: let haftmann@26064: fun add_stmts names (gr, (maxidx, idx_tab)) = if exists ((can o Graph.get_node) gr) names haftmann@26064: then (gr, (maxidx, idx_tab)) haftmann@26064: else (gr, (maxidx, idx_tab)) haftmann@27103: |> compile_stmts (map (fn name => ((name, Graph.get_node program name), haftmann@27103: Graph.imm_succs program name)) names); haftmann@27103: in fold_rev add_stmts (Graph.strong_conn program) end; haftmann@24155: haftmann@25944: haftmann@25944: (** evaluation **) haftmann@25944: haftmann@25944: (* term evaluation *) haftmann@25944: haftmann@25924: fun eval_term gr deps ((vs, ty), t) = haftmann@25924: let haftmann@25924: val frees = CodeThingol.fold_unbound_varnames (insert (op =)) t [] haftmann@25924: val frees' = map (fn v => Free (v, [])) frees; haftmann@25924: val dict_frees = maps (fn (v, sort) => map_index (curry DFree v o fst) sort) vs; haftmann@25924: in haftmann@25935: ("", (vs, [(map IVar frees, t)])) haftmann@25924: |> singleton (compile_eqnss gr deps) haftmann@25924: |> snd haftmann@25924: |> (fn t => apps t (rev (dict_frees @ frees'))) haftmann@25924: end; haftmann@24155: haftmann@24839: (* reification *) haftmann@24155: haftmann@26064: fun term_of_univ thy idx_tab t = haftmann@24155: let haftmann@25101: fun take_until f [] = [] haftmann@25101: | take_until f (x::xs) = if f x then [] else x :: take_until f xs; haftmann@26064: fun is_dict (Const (idx, _)) = haftmann@26064: let haftmann@26064: val c = the (Inttab.lookup idx_tab idx); haftmann@26064: in haftmann@26064: (is_some o CodeName.class_rev thy) c haftmann@26064: orelse (is_some o CodeName.classrel_rev thy) c haftmann@26064: orelse (is_some o CodeName.instance_rev thy) c haftmann@26064: end haftmann@25101: | is_dict (DFree _) = true haftmann@25101: | is_dict _ = false; haftmann@24155: fun of_apps bounds (t, ts) = haftmann@24155: fold_map (of_univ bounds) ts haftmann@24155: #>> (fn ts' => list_comb (t, rev ts')) haftmann@26064: and of_univ bounds (Const (idx, ts)) typidx = haftmann@24155: let haftmann@25101: val ts' = take_until is_dict ts; haftmann@26064: val c = (the o CodeName.const_rev thy o the o Inttab.lookup idx_tab) idx; haftmann@26970: val (_, T) = Code.default_typ thy c; haftmann@26739: val T' = map_type_tvar (fn ((v, i), S) => TypeInfer.param (typidx + i) (v, [])) T; haftmann@24155: val typidx' = typidx + maxidx_of_typ T' + 1; haftmann@25101: in of_apps bounds (Term.Const (c, T'), ts') typidx' end haftmann@24155: | of_univ bounds (Free (name, ts)) typidx = haftmann@24155: of_apps bounds (Term.Free (name, dummyT), ts) typidx haftmann@27499: | of_univ bounds (BVar (n, ts)) typidx = haftmann@27499: of_apps bounds (Bound (bounds - n - 1), ts) typidx haftmann@24155: | of_univ bounds (t as Abs _) typidx = haftmann@24155: typidx haftmann@25924: |> of_univ (bounds + 1) (apps t [BVar (bounds, [])]) haftmann@24155: |-> (fn t' => pair (Term.Abs ("u", dummyT, t'))) haftmann@24155: in of_univ 0 t 0 |> fst end; haftmann@24155: haftmann@25101: (* function store *) haftmann@25101: haftmann@25101: structure Nbe_Functions = CodeDataFun haftmann@25101: ( haftmann@26064: type T = (Univ option * int) Graph.T * (int * string Inttab.table); haftmann@26064: val empty = (Graph.empty, (0, Inttab.empty)); haftmann@27609: fun purge thy cs (gr, (maxidx, idx_tab)) = haftmann@27609: let haftmann@27609: val cs_exisiting = haftmann@27609: map_filter (CodeName.const_rev thy) (Graph.keys gr); haftmann@27609: val dels = (Graph.all_preds gr haftmann@27609: o map (CodeName.const thy) haftmann@27609: o filter (member (op =) cs_exisiting) haftmann@27609: ) cs; haftmann@27609: in (Graph.del_nodes dels gr, (maxidx, idx_tab)) end; haftmann@25101: ); haftmann@25101: haftmann@25101: (* compilation, evaluation and reification *) haftmann@25101: haftmann@27103: fun compile_eval thy program vs_ty_t deps = haftmann@26064: let haftmann@27103: val (gr, (_, idx_tab)) = Nbe_Functions.change thy (ensure_stmts program); haftmann@26064: in haftmann@26064: vs_ty_t haftmann@26064: |> eval_term gr deps haftmann@26064: |> term_of_univ thy idx_tab haftmann@26064: end; haftmann@25101: haftmann@24155: (* evaluation with type reconstruction *) haftmann@24155: haftmann@27103: fun eval thy t program vs_ty_t deps = haftmann@24155: let haftmann@26747: fun subst_const f = map_aterms (fn t as Term.Const (c, ty) => Term.Const (f c, ty) haftmann@26747: | t => t); haftmann@26747: val subst_triv_consts = subst_const (CodeUnit.resubst_alias thy); haftmann@24347: val ty = type_of t; haftmann@26739: val type_free = AList.lookup (op =) haftmann@26739: (map (fn (s, T) => (s, Term.Free (s, T))) (Term.add_frees t [])); haftmann@26739: val type_frees = Term.map_aterms haftmann@26739: (fn (t as Term.Free (s, _)) => the_default t (type_free s) | t => t); wenzelm@27264: fun type_infer t = wenzelm@27264: singleton (TypeInfer.infer_types (Syntax.pp_global thy) (Sign.tsig_of thy) I wenzelm@27264: (try (Type.strip_sorts o Sign.the_const_type thy)) (K NONE) Name.context 0) wenzelm@27264: (TypeInfer.constrain ty t); haftmann@24155: fun check_tvars t = if null (Term.term_tvars t) then t else haftmann@24155: error ("Illegal schematic type variables in normalized term: " wenzelm@26952: ^ setmp show_types true (Syntax.string_of_term_global thy) t); wenzelm@26952: val string_of_term = setmp show_types true (Syntax.string_of_term_global thy); haftmann@24155: in haftmann@27103: compile_eval thy program vs_ty_t deps haftmann@25167: |> tracing (fn t => "Normalized:\n" ^ string_of_term t) haftmann@26747: |> subst_triv_consts haftmann@26739: |> type_frees haftmann@25167: |> tracing (fn t => "Vars typed:\n" ^ string_of_term t) haftmann@26739: |> type_infer haftmann@25167: |> tracing (fn t => "Types inferred:\n" ^ string_of_term t) haftmann@26739: |> check_tvars haftmann@25101: |> tracing (fn t => "---\n") haftmann@24155: end; haftmann@24155: haftmann@24155: (* evaluation oracle *) haftmann@24155: haftmann@27103: exception Norm of term * CodeThingol.program haftmann@24381: * (CodeThingol.typscheme * CodeThingol.iterm) * string list; haftmann@24155: haftmann@27103: fun norm_oracle (thy, Norm (t, program, vs_ty_t, deps)) = haftmann@27103: Logic.mk_equals (t, eval thy t program vs_ty_t deps); haftmann@24155: haftmann@27103: fun norm_invoke thy t program vs_ty_t deps = haftmann@27103: Thm.invoke_oracle_i thy "HOL.norm" (thy, Norm (t, program, vs_ty_t, deps)); haftmann@24283: (*FIXME get rid of hardwired theory name*) haftmann@24155: haftmann@26747: fun add_triv_classes thy = haftmann@26747: let haftmann@26747: val inters = curry (Sorts.inter_sort (Sign.classes_of thy)) haftmann@26747: (CodeUnit.triv_classes thy); haftmann@26747: fun map_sorts f = (map_types o map_atyps) haftmann@26747: (fn TVar (v, sort) => TVar (v, f sort) haftmann@26747: | TFree (v, sort) => TFree (v, f sort)); haftmann@26747: in map_sorts inters end; haftmann@26747: haftmann@24839: fun norm_conv ct = haftmann@24155: let haftmann@24155: val thy = Thm.theory_of_cterm ct; haftmann@27103: fun evaluator' t program vs_ty_t deps = norm_invoke thy t program vs_ty_t deps; haftmann@26739: fun evaluator t = (add_triv_classes thy t, evaluator' t); haftmann@27103: in CodeThingol.eval_conv thy evaluator ct end; haftmann@24155: haftmann@26739: fun norm_term thy t = haftmann@24839: let haftmann@27103: fun evaluator' t program vs_ty_t deps = eval thy t program vs_ty_t deps; haftmann@26739: fun evaluator t = (add_triv_classes thy t, evaluator' t); haftmann@27103: in (Code.postprocess_term thy o CodeThingol.eval_term thy evaluator) t end; haftmann@24839: haftmann@24155: (* evaluation command *) haftmann@24155: haftmann@24155: fun norm_print_term ctxt modes t = haftmann@24155: let haftmann@24155: val thy = ProofContext.theory_of ctxt; haftmann@24839: val t' = norm_term thy t; haftmann@24839: val ty' = Term.type_of t'; wenzelm@26952: val ctxt' = Variable.auto_fixes t ctxt; wenzelm@24634: val p = PrintMode.with_modes modes (fn () => wenzelm@26952: Pretty.block [Pretty.quote (Syntax.pretty_term ctxt' t'), Pretty.fbrk, wenzelm@26952: Pretty.str "::", Pretty.brk 1, Pretty.quote (Syntax.pretty_typ ctxt' ty')]) (); haftmann@24155: in Pretty.writeln p end; haftmann@24155: haftmann@24155: haftmann@24155: (** Isar setup **) haftmann@24155: wenzelm@24508: fun norm_print_term_cmd (modes, s) state = haftmann@24155: let val ctxt = Toplevel.context_of state wenzelm@24508: in norm_print_term ctxt modes (Syntax.read_term ctxt s) end; haftmann@24155: haftmann@26747: val setup = Theory.add_oracle ("norm", norm_oracle); haftmann@24155: haftmann@24155: local structure P = OuterParse and K = OuterKeyword in haftmann@24155: haftmann@24155: val opt_modes = Scan.optional (P.$$$ "(" |-- P.!!! (Scan.repeat1 P.xname --| P.$$$ ")")) []; haftmann@24155: wenzelm@24867: val _ = haftmann@24155: OuterSyntax.improper_command "normal_form" "normalize term by evaluation" K.diag haftmann@24155: (opt_modes -- P.typ >> (Toplevel.keep o norm_print_term_cmd)); haftmann@24155: haftmann@24155: end; haftmann@24155: haftmann@24155: end;