haftmann@24219: (* Title: Tools/code/code_package.ML haftmann@24219: ID: $Id$ haftmann@24219: Author: Florian Haftmann, TU Muenchen haftmann@24219: haftmann@24219: Code generator translation kernel. Code generator Isar setup. haftmann@24219: *) haftmann@24219: haftmann@24219: signature CODE_PACKAGE = haftmann@24219: sig haftmann@24219: val eval_conv: theory haftmann@24381: -> (CodeThingol.code -> CodeThingol.typscheme * CodeThingol.iterm haftmann@24381: -> string list -> cterm -> thm) haftmann@24283: -> cterm -> thm; haftmann@24283: val eval_term: theory haftmann@24381: -> (CodeThingol.code -> CodeThingol.typscheme * CodeThingol.iterm haftmann@24835: -> string list -> term -> 'a) haftmann@24835: -> term -> 'a; haftmann@24585: val satisfies_ref: (unit -> bool) option ref; haftmann@24835: val satisfies: theory -> term -> string list -> bool; haftmann@24621: val eval_invoke: theory -> (string * (unit -> 'a) option ref) -> CodeThingol.code haftmann@24621: -> CodeThingol.iterm * CodeThingol.itype -> string list -> 'a; haftmann@24219: val codegen_command: theory -> string -> unit; haftmann@24219: end; haftmann@24219: haftmann@24219: structure CodePackage : CODE_PACKAGE = haftmann@24219: struct haftmann@24219: haftmann@24219: open BasicCodeThingol; haftmann@24219: haftmann@24844: (** code theorems **) haftmann@24219: haftmann@24283: fun code_depgr thy [] = CodeFuncgr.make thy [] haftmann@24219: | code_depgr thy consts = haftmann@24219: let haftmann@24283: val gr = CodeFuncgr.make thy consts; haftmann@24423: val select = Graph.all_succs gr consts; haftmann@24219: in haftmann@24423: gr haftmann@24423: |> Graph.subgraph (member (op =) select) haftmann@24423: |> Graph.map_nodes ((apsnd o map) (Conv.fconv_rule (Class.overload thy))) haftmann@24219: end; haftmann@24219: haftmann@24219: fun code_thms thy = haftmann@24219: Pretty.writeln o CodeFuncgr.pretty thy o code_depgr thy; haftmann@24219: haftmann@24219: fun code_deps thy consts = haftmann@24219: let haftmann@24219: val gr = code_depgr thy consts; haftmann@24219: fun mk_entry (const, (_, (_, parents))) = haftmann@24219: let haftmann@24219: val name = CodeUnit.string_of_const thy const; haftmann@24219: val nameparents = map (CodeUnit.string_of_const thy) parents; haftmann@24219: in { name = name, ID = name, dir = "", unfold = true, haftmann@24219: path = "", parents = nameparents } haftmann@24219: end; haftmann@24423: val prgr = Graph.fold ((fn x => fn xs => xs @ [x]) o mk_entry) gr []; haftmann@24219: in Present.display_graph prgr end; haftmann@24219: haftmann@24844: haftmann@24844: (** code translation **) haftmann@24844: haftmann@24219: structure Program = CodeDataFun haftmann@24219: ( haftmann@24219: type T = CodeThingol.code; haftmann@24219: val empty = CodeThingol.empty_code; haftmann@24219: fun merge _ = CodeThingol.merge_code; haftmann@24219: fun purge _ NONE _ = CodeThingol.empty_code haftmann@24219: | purge NONE _ _ = CodeThingol.empty_code haftmann@24219: | purge (SOME thy) (SOME cs) code = haftmann@24219: let haftmann@24219: val cs_exisiting = haftmann@24219: map_filter (CodeName.const_rev thy) (Graph.keys code); haftmann@24219: val dels = (Graph.all_preds code haftmann@24219: o map (CodeName.const thy) haftmann@24423: o filter (member (op =) cs_exisiting) haftmann@24219: ) cs; haftmann@24219: in Graph.del_nodes dels code end; haftmann@24219: ); haftmann@24219: haftmann@24219: haftmann@24219: (* translation kernel *) haftmann@24219: haftmann@24283: val value_name = "Isabelle_Eval.EVAL.EVAL"; haftmann@24283: haftmann@24283: fun ensure_def thy = CodeThingol.ensure_def haftmann@24283: (fn s => if s = value_name then "" else CodeName.labelled_name thy s); haftmann@24219: haftmann@24811: exception CONSTRAIN of (string * typ) * typ; haftmann@24811: haftmann@24219: fun ensure_def_class thy (algbr as ((_, algebra), _)) funcgr class = haftmann@24219: let haftmann@24219: val superclasses = (Sorts.certify_sort algebra o Sorts.super_classes algebra) class; haftmann@24219: val (v, cs) = AxClass.params_of_class thy class; haftmann@24219: val class' = CodeName.class thy class; haftmann@24219: val defgen_class = haftmann@24811: fold_map (fn superclass => ensure_def_class thy algbr funcgr superclass haftmann@24811: ##>> ensure_def_classrel thy algbr funcgr (class, superclass)) superclasses haftmann@24811: ##>> fold_map (fn (c, ty) => ensure_def_const thy algbr funcgr c haftmann@24811: ##>> exprgen_typ thy algbr funcgr ty) cs haftmann@24811: #>> (fn info => CodeThingol.Class (unprefix "'" v, info)) haftmann@24219: in haftmann@24811: ensure_def thy defgen_class class' haftmann@24219: #> pair class' haftmann@24219: end haftmann@24219: and ensure_def_classrel thy algbr funcgr (subclass, superclass) = haftmann@24811: let haftmann@24811: val classrel' = CodeName.classrel thy (subclass, superclass); haftmann@24811: val defgen_classrel = haftmann@24811: ensure_def_class thy algbr funcgr subclass haftmann@24811: ##>> ensure_def_class thy algbr funcgr superclass haftmann@24811: #>> CodeThingol.Classrel; haftmann@24811: in haftmann@24811: ensure_def thy defgen_classrel classrel' haftmann@24811: #> pair classrel' haftmann@24811: end haftmann@24250: and ensure_def_tyco thy algbr funcgr "fun" = haftmann@24250: pair "fun" haftmann@24250: | ensure_def_tyco thy algbr funcgr tyco = haftmann@24219: let haftmann@24250: val defgen_datatype = haftmann@24219: let haftmann@24219: val (vs, cos) = Code.get_datatype thy tyco; haftmann@24219: in haftmann@24250: fold_map (exprgen_tyvar_sort thy algbr funcgr) vs haftmann@24250: ##>> fold_map (fn (c, tys) => haftmann@24811: ensure_def_const thy algbr funcgr c haftmann@24811: ##>> fold_map (exprgen_typ thy algbr funcgr) tys) cos haftmann@24811: #>> CodeThingol.Datatype haftmann@24219: end; haftmann@24219: val tyco' = CodeName.tyco thy tyco; haftmann@24219: in haftmann@24811: ensure_def thy defgen_datatype tyco' haftmann@24250: #> pair tyco' haftmann@24219: end haftmann@24585: and exprgen_tyvar_sort thy (algbr as ((proj_sort, _), _)) funcgr (v, sort) = haftmann@24585: fold_map (ensure_def_class thy algbr funcgr) (proj_sort sort) haftmann@24585: #>> (fn sort => (unprefix "'" v, sort)) haftmann@24585: and exprgen_typ thy algbr funcgr (TFree vs) = haftmann@24585: exprgen_tyvar_sort thy algbr funcgr vs haftmann@24585: #>> (fn (v, sort) => ITyVar v) haftmann@24585: | exprgen_typ thy algbr funcgr (Type (tyco, tys)) = haftmann@24585: ensure_def_tyco thy algbr funcgr tyco haftmann@24585: ##>> fold_map (exprgen_typ thy algbr funcgr) tys haftmann@24811: #>> (fn (tyco, tys) => tyco `%% tys) haftmann@24811: and exprgen_dicts thy (algbr as ((proj_sort, algebra), consts)) funcgr (ty_ctxt, sort_decl) = haftmann@24219: let haftmann@24219: val pp = Sign.pp thy; haftmann@24219: datatype typarg = haftmann@24219: Global of (class * string) * typarg list list haftmann@24219: | Local of (class * class) list * (string * (int * sort)); haftmann@24219: fun class_relation (Global ((_, tyco), yss), _) class = haftmann@24219: Global ((class, tyco), yss) haftmann@24219: | class_relation (Local (classrels, v), subclass) superclass = haftmann@24219: Local ((subclass, superclass) :: classrels, v); haftmann@24219: fun type_constructor tyco yss class = haftmann@24219: Global ((class, tyco), (map o map) fst yss); haftmann@24219: fun type_variable (TFree (v, sort)) = haftmann@24219: let haftmann@24219: val sort' = proj_sort sort; haftmann@24219: in map_index (fn (n, class) => (Local ([], (v, (n, sort'))), class)) sort' end; haftmann@24219: val typargs = Sorts.of_sort_derivation pp algebra haftmann@24219: {class_relation = class_relation, type_constructor = type_constructor, haftmann@24219: type_variable = type_variable} haftmann@24219: (ty_ctxt, proj_sort sort_decl); haftmann@24219: fun mk_dict (Global (inst, yss)) = haftmann@24219: ensure_def_inst thy algbr funcgr inst haftmann@24219: ##>> (fold_map o fold_map) mk_dict yss haftmann@24219: #>> (fn (inst, dss) => DictConst (inst, dss)) haftmann@24219: | mk_dict (Local (classrels, (v, (k, sort)))) = haftmann@24219: fold_map (ensure_def_classrel thy algbr funcgr) classrels haftmann@24219: #>> (fn classrels => DictVar (classrels, (unprefix "'" v, (k, length sort)))) haftmann@24219: in haftmann@24219: fold_map mk_dict typargs haftmann@24219: end haftmann@24219: and exprgen_dict_parms thy (algbr as (_, consts)) funcgr (c, ty_ctxt) = haftmann@24219: let haftmann@24423: val ty_decl = Consts.the_declaration consts c; haftmann@24423: val (tys, tys_decl) = pairself (curry (Consts.typargs consts) c) (ty_ctxt, ty_decl); haftmann@24219: val sorts = map (snd o dest_TVar) tys_decl; haftmann@24219: in haftmann@24219: fold_map (exprgen_dicts thy algbr funcgr) (tys ~~ sorts) haftmann@24219: end haftmann@24585: and exprgen_eq thy algbr funcgr thm = haftmann@24585: let haftmann@24585: val (args, rhs) = (apfst (snd o strip_comb) o Logic.dest_equals haftmann@24585: o Logic.unvarify o prop_of) thm; haftmann@24585: in haftmann@24585: fold_map (exprgen_term thy algbr funcgr) args haftmann@24585: ##>> exprgen_term thy algbr funcgr rhs haftmann@24585: #>> rpair thm haftmann@24585: end haftmann@24585: and ensure_def_inst thy (algbr as ((_, algebra), _)) funcgr (class, tyco) = haftmann@24219: let haftmann@24219: val superclasses = (Sorts.certify_sort algebra o Sorts.super_classes algebra) class; haftmann@24835: val (var, classparams) = try (AxClass.params_of_class thy) class |> the_default ("'a", []) haftmann@24585: val vs = Name.names Name.context "'a" (Sorts.mg_domain algebra tyco [class]); haftmann@24585: val sorts' = Sorts.mg_domain (Sign.classes_of thy) tyco [class]; haftmann@24585: val vs' = map2 (fn (v, sort1) => fn sort2 => (v, haftmann@24585: Sorts.inter_sort (Sign.classes_of thy) (sort1, sort2))) vs sorts'; haftmann@24219: val arity_typ = Type (tyco, map TFree vs); haftmann@24585: val arity_typ' = Type (tyco, map (fn (v, sort) => TVar ((v, 0), sort)) vs'); haftmann@24835: fun exprgen_superarity superclass = haftmann@24585: ensure_def_class thy algbr funcgr superclass haftmann@24585: ##>> ensure_def_classrel thy algbr funcgr (class, superclass) haftmann@24585: ##>> exprgen_dicts thy algbr funcgr (arity_typ, [superclass]) haftmann@24585: #>> (fn ((superclass, classrel), [DictConst (inst, dss)]) => haftmann@24219: (superclass, (classrel, (inst, dss)))); haftmann@24835: fun exprgen_classparam_inst (c, ty) = haftmann@24585: let haftmann@24585: val c_inst = Const (c, map_type_tfree (K arity_typ') ty); haftmann@24585: val thm = Class.unoverload thy (Thm.cterm_of thy c_inst); haftmann@24585: val c_ty = (apsnd Logic.unvarifyT o dest_Const o snd haftmann@24585: o Logic.dest_equals o Thm.prop_of) thm; haftmann@24585: in haftmann@24585: ensure_def_const thy algbr funcgr c haftmann@24585: ##>> exprgen_const thy algbr funcgr c_ty haftmann@24585: #>> (fn (c, IConst c_inst) => ((c, c_inst), thm)) haftmann@24585: end; haftmann@24585: val defgen_inst = haftmann@24585: ensure_def_class thy algbr funcgr class haftmann@24585: ##>> ensure_def_tyco thy algbr funcgr tyco haftmann@24585: ##>> fold_map (exprgen_tyvar_sort thy algbr funcgr) vs haftmann@24835: ##>> fold_map exprgen_superarity superclasses haftmann@24835: ##>> fold_map exprgen_classparam_inst classparams haftmann@24835: #>> (fn ((((class, tyco), arity), superarities), classparams) => haftmann@24835: CodeThingol.Classinst ((class, (tyco, arity)), (superarities, classparams))); haftmann@24219: val inst = CodeName.instance thy (class, tyco); haftmann@24219: in haftmann@24811: ensure_def thy defgen_inst inst haftmann@24585: #> pair inst haftmann@24219: end haftmann@24423: and ensure_def_const thy (algbr as (_, consts)) funcgr c = haftmann@24219: let haftmann@24423: val c' = CodeName.const thy c; haftmann@24585: fun defgen_datatypecons tyco = haftmann@24585: ensure_def_tyco thy algbr funcgr tyco haftmann@24811: #>> K (CodeThingol.Datatypecons c'); haftmann@24835: fun defgen_classparam class = haftmann@24585: ensure_def_class thy algbr funcgr class haftmann@24835: #>> K (CodeThingol.Classparam c'); haftmann@24219: fun defgen_fun trns = haftmann@24219: let haftmann@24423: val raw_thms = CodeFuncgr.funcs funcgr c; haftmann@24423: val ty = (Logic.unvarifyT o CodeFuncgr.typ funcgr) c; haftmann@24585: val vs = (map dest_TFree o Consts.typargs consts) (c, ty); haftmann@24219: val thms = if (null o Term.typ_tfrees) ty orelse (null o fst o strip_type) ty haftmann@24219: then raw_thms haftmann@24219: else map (CodeUnit.expand_eta 1) raw_thms; haftmann@24219: in haftmann@24219: trns haftmann@24381: |> fold_map (exprgen_tyvar_sort thy algbr funcgr) vs haftmann@24219: ||>> exprgen_typ thy algbr funcgr ty haftmann@24585: ||>> fold_map (exprgen_eq thy algbr funcgr) thms haftmann@24381: |>> (fn ((vs, ty), eqs) => CodeThingol.Fun ((vs, ty), eqs)) haftmann@24219: end; haftmann@24585: val defgen = case Code.get_datatype_of_constr thy c haftmann@24585: of SOME tyco => defgen_datatypecons tyco haftmann@24585: | NONE => (case AxClass.class_of_param thy c haftmann@24835: of SOME class => defgen_classparam class haftmann@24585: | NONE => defgen_fun) haftmann@24219: in haftmann@24811: ensure_def thy defgen c' haftmann@24250: #> pair c' haftmann@24219: end haftmann@24585: and exprgen_term thy algbr funcgr (Const (c, ty)) = haftmann@24835: exprgen_app thy algbr funcgr ((c, ty), []) haftmann@24585: | exprgen_term thy algbr funcgr (Free (v, _)) = haftmann@24585: pair (IVar v) haftmann@24585: | exprgen_term thy algbr funcgr (Abs (abs as (_, ty, _))) = haftmann@24219: let haftmann@24585: val (v, t) = Syntax.variant_abs abs; haftmann@24219: in haftmann@24585: exprgen_typ thy algbr funcgr ty haftmann@24585: ##>> exprgen_term thy algbr funcgr t haftmann@24585: #>> (fn (ty, t) => (v, ty) `|-> t) haftmann@24219: end haftmann@24585: | exprgen_term thy algbr funcgr (t as _ $ _) = haftmann@24219: case strip_comb t haftmann@24219: of (Const (c, ty), ts) => haftmann@24835: exprgen_app thy algbr funcgr ((c, ty), ts) haftmann@24219: | (t', ts) => haftmann@24585: exprgen_term thy algbr funcgr t' haftmann@24585: ##>> fold_map (exprgen_term thy algbr funcgr) ts haftmann@24585: #>> (fn (t, ts) => t `$$ ts) haftmann@24585: and exprgen_const thy algbr funcgr (c, ty) = haftmann@24585: ensure_def_const thy algbr funcgr c haftmann@24585: ##>> exprgen_dict_parms thy algbr funcgr (c, ty) haftmann@24585: ##>> fold_map (exprgen_typ thy algbr funcgr) ((fst o Term.strip_type) ty) haftmann@24585: (*##>> exprgen_typ thy algbr funcgr ((snd o Term.strip_type) ty)*) haftmann@24585: #>> (fn ((c, iss), tys) => IConst (c, (iss, tys))) haftmann@24585: and exprgen_app_default thy algbr funcgr (c_ty, ts) = haftmann@24585: exprgen_const thy algbr funcgr c_ty haftmann@24585: ##>> fold_map (exprgen_term thy algbr funcgr) ts haftmann@24585: #>> (fn (t, ts) => t `$$ ts) haftmann@24844: and exprgen_case thy algbr funcgr n cases (app as ((c, ty), ts)) = haftmann@24844: let haftmann@24844: (*FIXME rework this code*) haftmann@24844: val (all_tys, _) = strip_type ty; haftmann@24844: val (tys, _) = chop (1 + (if null cases then 1 else length cases)) all_tys; haftmann@24844: val st = nth ts n; haftmann@24844: val sty = nth tys n; haftmann@24844: fun is_undefined (Const (c, _)) = Code.is_undefined thy c haftmann@24844: | is_undefined _ = false; haftmann@24844: fun mk_case (co, n) t = haftmann@24844: let haftmann@24844: val (vs, body) = Term.strip_abs_eta n t; haftmann@24844: val selector = list_comb (Const (co, map snd vs ---> sty), map Free vs); haftmann@24844: in if is_undefined body then NONE else SOME (selector, body) end; haftmann@24844: val ds = case cases haftmann@24844: of [] => let val ([v_ty], body) = Term.strip_abs_eta 1 (the_single (nth_drop n ts)) haftmann@24844: in [(Free v_ty, body)] end haftmann@24844: | cases => map_filter (uncurry mk_case) (AList.make (CodeUnit.no_args thy) cases haftmann@24844: ~~ nth_drop n ts); haftmann@24844: in haftmann@24844: exprgen_term thy algbr funcgr st haftmann@24844: ##>> exprgen_typ thy algbr funcgr sty haftmann@24844: ##>> fold_map (fn (pat, body) => exprgen_term thy algbr funcgr pat haftmann@24844: ##>> exprgen_term thy algbr funcgr body) ds haftmann@24844: ##>> exprgen_app_default thy algbr funcgr app haftmann@24844: #>> (fn (((st, sty), ds), t0) => ICase (((st, sty), ds), t0)) haftmann@24844: end haftmann@24844: and exprgen_app thy algbr funcgr ((c, ty), ts) = case Code.get_case_data thy c haftmann@24844: of SOME (n, cases) => let val i = 1 + (if null cases then 1 else length cases) in haftmann@24844: if length ts < i then haftmann@24844: let haftmann@24844: val k = length ts; haftmann@24844: val tys = (curry Library.take (i - k) o curry Library.drop k o fst o strip_type) ty; haftmann@24844: val ctxt = (fold o fold_aterms) haftmann@24844: (fn Free (v, _) => Name.declare v | _ => I) ts Name.context; haftmann@24844: val vs = Name.names ctxt "a" tys; haftmann@24844: in haftmann@24844: fold_map (exprgen_typ thy algbr funcgr) tys haftmann@24844: ##>> exprgen_case thy algbr funcgr n cases ((c, ty), ts @ map Free vs) haftmann@24844: #>> (fn (tys, t) => map2 (fn (v, _) => pair v) vs tys `|--> t) haftmann@24844: end haftmann@24844: else if length ts > i then haftmann@24844: exprgen_case thy algbr funcgr n cases ((c, ty), Library.take (i, ts)) haftmann@24844: ##>> fold_map (exprgen_term thy algbr funcgr) (Library.drop (i, ts)) haftmann@24844: #>> (fn (t, ts) => t `$$ ts) haftmann@24844: else haftmann@24844: exprgen_case thy algbr funcgr n cases ((c, ty), ts) haftmann@24844: end haftmann@24844: | NONE => exprgen_app_default thy algbr funcgr ((c, ty), ts); haftmann@24219: haftmann@24219: haftmann@24219: (* entrance points into translation kernel *) haftmann@24219: haftmann@24219: fun ensure_def_const' thy algbr funcgr c trns = haftmann@24219: ensure_def_const thy algbr funcgr c trns haftmann@24219: handle CONSTRAIN ((c, ty), ty_decl) => error ( haftmann@24219: "Constant " ^ c ^ " with most general type\n" haftmann@24219: ^ CodeUnit.string_of_typ thy ty haftmann@24219: ^ "\noccurs with type\n" haftmann@24219: ^ CodeUnit.string_of_typ thy ty_decl); haftmann@24219: haftmann@24219: fun perhaps_def_const thy algbr funcgr c trns = haftmann@24219: case try (ensure_def_const thy algbr funcgr c) trns haftmann@24219: of SOME (c, trns) => (SOME c, trns) haftmann@24219: | NONE => (NONE, trns); haftmann@24219: haftmann@24219: fun exprgen_term' thy algbr funcgr t trns = haftmann@24219: exprgen_term thy algbr funcgr t trns haftmann@24219: handle CONSTRAIN ((c, ty), ty_decl) => error ("In term " ^ (quote o Sign.string_of_term thy) t haftmann@24219: ^ ",\nconstant " ^ c ^ " with most general type\n" haftmann@24219: ^ CodeUnit.string_of_typ thy ty haftmann@24219: ^ "\noccurs with type\n" haftmann@24219: ^ CodeUnit.string_of_typ thy ty_decl); haftmann@24219: haftmann@24219: haftmann@24219: (** code generation interfaces **) haftmann@24219: haftmann@24219: (* generic generation combinators *) haftmann@24219: haftmann@24219: fun generate thy funcgr gen it = haftmann@24219: let haftmann@24219: val naming = NameSpace.qualified_names NameSpace.default_naming; haftmann@24219: val consttab = Consts.empty wenzelm@24770: |> fold (fn c => Consts.declare true naming [] (c, CodeFuncgr.typ funcgr c)) haftmann@24423: (CodeFuncgr.all funcgr); haftmann@24219: val algbr = (Code.operational_algebra thy, consttab); haftmann@24219: in haftmann@24219: Program.change_yield thy haftmann@24423: (CodeThingol.start_transact (gen thy algbr funcgr it)) haftmann@24283: |> fst haftmann@24219: end; haftmann@24219: haftmann@24436: fun code thy permissive cs seris = haftmann@24436: let haftmann@24436: val code = Program.get thy; haftmann@24436: val seris' = map (fn (((target, module), file), args) => haftmann@24436: CodeTarget.get_serializer thy target permissive module file args haftmann@24835: CodeName.labelled_name cs) seris; haftmann@24436: in (map (fn f => f code) seris' : unit list; ()) end; haftmann@24436: haftmann@24835: fun raw_eval evaluate term_of thy g = haftmann@24250: let haftmann@24250: val value_name = "Isabelle_Eval.EVAL.EVAL"; haftmann@24250: fun ensure_eval thy algbr funcgr t = haftmann@24250: let haftmann@24381: val ty = fastype_of t; haftmann@24436: val vs = fold_term_types (K (fold_atyps (insert (eq_fst op =) haftmann@24436: o dest_TFree))) t []; haftmann@24250: val defgen_eval = haftmann@24381: fold_map (exprgen_tyvar_sort thy algbr funcgr) vs haftmann@24381: ##>> exprgen_typ thy algbr funcgr ty haftmann@24381: ##>> exprgen_term' thy algbr funcgr t haftmann@24585: #>> (fn ((vs, ty), t) => CodeThingol.Fun ((vs, ty), [(([], t), Drule.dummy_thm)])); haftmann@24283: fun result (dep, code) = haftmann@24283: let haftmann@24585: val CodeThingol.Fun ((vs, ty), [(([], t), _)]) = Graph.get_node code value_name; haftmann@24283: val deps = Graph.imm_succs code value_name; haftmann@24283: val code' = Graph.del_nodes [value_name] code; haftmann@24283: val code'' = CodeThingol.project_code false [] (SOME deps) code'; haftmann@24381: in ((code'', ((vs, ty), t), deps), (dep, code')) end; haftmann@24250: in haftmann@24811: ensure_def thy defgen_eval value_name haftmann@24283: #> result haftmann@24250: end; haftmann@24283: fun h funcgr ct = haftmann@24219: let haftmann@24835: val (code, vs_ty_t, deps) = generate thy funcgr ensure_eval (term_of ct); haftmann@24381: in g code vs_ty_t deps ct end; haftmann@24835: in evaluate thy h end; haftmann@24219: haftmann@24835: fun eval_conv thy = raw_eval CodeFuncgr.eval_conv Thm.term_of thy; haftmann@24835: fun eval_term thy = raw_eval CodeFuncgr.eval_term I thy; haftmann@24219: haftmann@24835: fun eval_invoke thy = CodeTarget.eval_invoke thy CodeName.labelled_name; haftmann@24621: haftmann@24585: val satisfies_ref : (unit -> bool) option ref = ref NONE; haftmann@24219: haftmann@24835: fun satisfies thy t witnesses = haftmann@24283: let haftmann@24381: fun evl code ((vs, ty), t) deps ct = haftmann@24662: eval_invoke thy ("CodePackage.satisfies_ref", satisfies_ref) haftmann@24662: code (t, ty) witnesses; haftmann@24835: in eval_term thy evl t end; haftmann@24219: haftmann@24219: fun filter_generatable thy consts = haftmann@24219: let haftmann@24283: val (consts', funcgr) = CodeFuncgr.make_consts thy consts; haftmann@24283: val consts'' = generate thy funcgr (fold_map ooo perhaps_def_const) consts'; haftmann@24219: val consts''' = map_filter (fn (const, SOME _) => SOME const | (_, NONE) => NONE) haftmann@24219: (consts' ~~ consts''); haftmann@24219: in consts''' end; haftmann@24219: haftmann@24436: fun generate_const_exprs thy raw_cs = haftmann@24436: let haftmann@24436: val (perm1, cs) = CodeUnit.read_const_exprs thy haftmann@24436: (filter_generatable thy) raw_cs; haftmann@24436: val (perm2, cs') = case generate thy (CodeFuncgr.make thy cs) haftmann@24436: (fold_map ooo ensure_def_const') cs haftmann@24436: of [] => (true, NONE) haftmann@24436: | cs => (false, SOME cs); haftmann@24436: in (perm1 orelse perm2, cs') end; haftmann@24436: haftmann@24436: haftmann@24436: (** code properties **) haftmann@24436: haftmann@24436: fun mk_codeprops thy all_cs sel_cs = haftmann@24436: let haftmann@24436: fun select (thmref, thm) = case try (Drule.unvarify o Drule.zero_var_indexes) thm haftmann@24436: of NONE => NONE haftmann@24436: | SOME thm => let haftmann@24436: val t = (ObjectLogic.drop_judgment thy o Thm.prop_of) thm; haftmann@24436: val cs = fold_aterms (fn Const (c, ty) => haftmann@24436: cons (Class.unoverload_const thy (c, ty)) | _ => I) t []; haftmann@24436: in if exists (member (op =) sel_cs) cs haftmann@24436: andalso forall (member (op =) all_cs) cs haftmann@24436: then SOME (thmref, thm) else NONE end; haftmann@24436: fun mk_codeprop (thmref, thm) = haftmann@24436: let haftmann@24436: val t = ObjectLogic.drop_judgment thy (Thm.prop_of thm); haftmann@24436: val ty_judg = fastype_of t; haftmann@24436: val tfrees1 = fold_aterms (fn Const (c, ty) => haftmann@24436: Term.add_tfreesT ty | _ => I) t []; haftmann@24436: val vars = Term.add_frees t []; haftmann@24436: val tfrees2 = fold (Term.add_tfreesT o snd) vars []; haftmann@24436: val tfrees' = subtract (op =) tfrees2 tfrees1 |> map TFree; haftmann@24436: val ty = map Term.itselfT tfrees' @ map snd vars ---> ty_judg; haftmann@24436: val tfree_vars = map Logic.mk_type tfrees'; haftmann@24436: val c = PureThy.string_of_thmref thmref haftmann@24436: |> NameSpace.explode haftmann@24436: |> (fn [x] => [x] | (x::xs) => xs) haftmann@24436: |> space_implode "_" haftmann@24436: val propdef = (((c, ty), tfree_vars @ map Free vars), t); haftmann@24436: in if c = "" then NONE else SOME (thmref, propdef) end; haftmann@24436: in haftmann@24436: PureThy.thms_containing thy ([], []) haftmann@24436: |> maps PureThy.selections haftmann@24436: |> map_filter select haftmann@24436: |> map_filter mk_codeprop haftmann@24436: end; haftmann@24436: haftmann@24436: fun add_codeprops all_cs sel_cs thy = haftmann@24436: let haftmann@24436: val codeprops = mk_codeprops thy all_cs sel_cs; haftmann@24436: fun lift_name_yield f x = (Name.context, x) |> f ||> snd; haftmann@24436: fun add (thmref, (((raw_c, ty), ts), t)) (names, thy) = haftmann@24436: let haftmann@24436: val _ = warning ("Adding theorem " ^ quote (PureThy.string_of_thmref thmref) haftmann@24436: ^ " as code property " ^ quote raw_c); haftmann@24436: val ([raw_c'], names') = Name.variants [raw_c] names; haftmann@24436: in haftmann@24436: thy haftmann@24436: |> PureThy.simple_def ("", []) (((raw_c', ty, Syntax.NoSyn), ts), t) haftmann@24436: ||> pair names' haftmann@24436: end; haftmann@24436: in haftmann@24436: thy haftmann@24436: |> Sign.sticky_prefix "codeprop" haftmann@24436: |> lift_name_yield (fold_map add codeprops) haftmann@24436: ||> Sign.restore_naming thy haftmann@24621: |-> (fn c_thms => fold (Code.add_func o snd) c_thms #> pair c_thms) haftmann@24436: end; haftmann@24436: haftmann@24436: haftmann@24219: haftmann@24219: (** toplevel interface and setup **) haftmann@24219: haftmann@24219: local haftmann@24219: haftmann@24219: structure P = OuterParse haftmann@24219: and K = OuterKeyword haftmann@24219: haftmann@24436: fun code_cmd raw_cs seris thy = haftmann@24219: let haftmann@24436: val (permissive, cs) = generate_const_exprs thy raw_cs; haftmann@24436: val _ = code thy permissive cs seris; haftmann@24436: in () end; haftmann@24219: haftmann@24219: fun code_thms_cmd thy = haftmann@24283: code_thms thy o snd o CodeUnit.read_const_exprs thy (fst o CodeFuncgr.make_consts thy); haftmann@24219: haftmann@24219: fun code_deps_cmd thy = haftmann@24283: code_deps thy o snd o CodeUnit.read_const_exprs thy (fst o CodeFuncgr.make_consts thy); haftmann@24219: haftmann@24436: fun code_props_cmd raw_cs seris thy = haftmann@24436: let haftmann@24436: val (_, all_cs) = generate_const_exprs thy ["*"]; haftmann@24436: val (permissive, cs) = generate_const_exprs thy raw_cs; haftmann@24436: val (c_thms, thy') = add_codeprops (map (the o CodeName.const_rev thy) (these all_cs)) haftmann@24436: (map (the o CodeName.const_rev thy) (these cs)) thy; haftmann@24436: val prop_cs = (filter_generatable thy' o map fst) c_thms; haftmann@24436: val _ = if null seris then [] else generate thy' (CodeFuncgr.make thy' prop_cs) haftmann@24436: (fold_map ooo ensure_def_const') prop_cs; haftmann@24436: val _ = if null seris then () else code thy' permissive haftmann@24436: (SOME (map (CodeName.const thy') prop_cs)) seris; haftmann@24436: in thy' end; haftmann@24436: haftmann@24250: val (inK, module_nameK, fileK) = ("in", "module_name", "file"); haftmann@24219: haftmann@24436: fun code_exprP cmd = haftmann@24219: (Scan.repeat P.term haftmann@24219: -- Scan.repeat (P.$$$ inK |-- P.name haftmann@24250: -- Scan.option (P.$$$ module_nameK |-- P.name) haftmann@24219: -- Scan.option (P.$$$ fileK |-- P.name) haftmann@24219: -- Scan.optional (P.$$$ "(" |-- P.arguments --| P.$$$ ")") [] haftmann@24436: ) >> (fn (raw_cs, seris) => cmd raw_cs seris)); haftmann@24219: wenzelm@24867: val _ = OuterSyntax.keywords [inK, module_nameK, fileK]; haftmann@24219: haftmann@24436: val (codeK, code_thmsK, code_depsK, code_propsK) = haftmann@24436: ("export_code", "code_thms", "code_deps", "code_props"); haftmann@24219: haftmann@24219: in haftmann@24219: wenzelm@24867: val _ = haftmann@24219: OuterSyntax.improper_command codeK "generate executable code for constants" haftmann@24436: K.diag (P.!!! (code_exprP code_cmd) >> (fn f => Toplevel.keep (f o Toplevel.theory_of))); haftmann@24219: haftmann@24219: fun codegen_command thy cmd = haftmann@24436: case Scan.read OuterLex.stopper (P.!!! (code_exprP code_cmd)) ((filter OuterLex.is_proper o OuterSyntax.scan) cmd) haftmann@24219: of SOME f => (writeln "Now generating code..."; f thy) haftmann@24219: | NONE => error ("Bad directive " ^ quote cmd); haftmann@24219: wenzelm@24867: val _ = haftmann@24219: OuterSyntax.improper_command code_thmsK "print system of defining equations for code" OuterKeyword.diag haftmann@24219: (Scan.repeat P.term haftmann@24219: >> (fn cs => Toplevel.no_timing o Toplevel.unknown_theory haftmann@24219: o Toplevel.keep ((fn thy => code_thms_cmd thy cs) o Toplevel.theory_of))); haftmann@24219: wenzelm@24867: val _ = haftmann@24219: OuterSyntax.improper_command code_depsK "visualize dependencies of defining equations for code" OuterKeyword.diag haftmann@24219: (Scan.repeat P.term haftmann@24219: >> (fn cs => Toplevel.no_timing o Toplevel.unknown_theory haftmann@24219: o Toplevel.keep ((fn thy => code_deps_cmd thy cs) o Toplevel.theory_of))); haftmann@24219: wenzelm@24867: val _ = haftmann@24436: OuterSyntax.command code_propsK "generate characteristic properties for executable constants" haftmann@24436: K.thy_decl (P.!!! (code_exprP code_props_cmd) >> Toplevel.theory); haftmann@24219: haftmann@24436: end; (*local*) haftmann@24436: haftmann@24436: end; (*struct*)