haftmann@24219: (* Title: Tools/code/code_funcgr.ML haftmann@24219: ID: $Id$ haftmann@24219: Author: Florian Haftmann, TU Muenchen haftmann@24219: haftmann@24219: Retrieving, normalizing and structuring defining equations in graph haftmann@24219: with explicit dependencies. haftmann@24219: *) haftmann@24219: haftmann@24219: signature CODE_FUNCGR = haftmann@24219: sig haftmann@24219: type T haftmann@24423: val funcs: T -> string -> thm list haftmann@26971: val typ: T -> string -> (string * sort) list * typ haftmann@24423: val all: T -> string list haftmann@24219: val pretty: theory -> T -> Pretty.T haftmann@24423: val make: theory -> string list -> T haftmann@26740: val eval_conv: theory -> (term -> term * (T -> term -> thm)) -> cterm -> thm haftmann@26740: val eval_term: theory -> (term -> term * (T -> term -> 'a)) -> term -> 'a haftmann@26971: val timing: bool ref haftmann@24219: end haftmann@24219: haftmann@24283: structure CodeFuncgr : CODE_FUNCGR = haftmann@24219: struct haftmann@24219: haftmann@24219: (** the graph type **) haftmann@24219: haftmann@26971: type T = (((string * sort) list * typ) * thm list) Graph.T; haftmann@24219: haftmann@24219: fun funcs funcgr = haftmann@24423: these o Option.map snd o try (Graph.get_node funcgr); haftmann@24219: haftmann@24219: fun typ funcgr = haftmann@24423: fst o Graph.get_node funcgr; haftmann@24219: haftmann@24423: fun all funcgr = Graph.keys funcgr; haftmann@24219: haftmann@24219: fun pretty thy funcgr = haftmann@24423: AList.make (snd o Graph.get_node funcgr) (Graph.keys funcgr) haftmann@24219: |> (map o apfst) (CodeUnit.string_of_const thy) haftmann@24219: |> sort (string_ord o pairself fst) haftmann@24219: |> map (fn (s, thms) => haftmann@24219: (Pretty.block o Pretty.fbreaks) ( haftmann@24219: Pretty.str s haftmann@24219: :: map Display.pretty_thm thms haftmann@24219: )) haftmann@24219: |> Pretty.chunks; haftmann@24219: haftmann@24219: haftmann@24219: (** generic combinators **) haftmann@24219: haftmann@24219: fun fold_consts f thms = haftmann@24219: thms haftmann@24219: |> maps (op :: o swap o apfst (snd o strip_comb) o Logic.dest_equals o Thm.plain_prop_of) haftmann@24219: |> (fold o fold_aterms) (fn Const c => f c | _ => I); haftmann@24219: haftmann@24219: fun consts_of (const, []) = [] haftmann@24423: | consts_of (const, thms as _ :: _) = haftmann@24219: let haftmann@24423: fun the_const (c, _) = if c = const then I else insert (op =) c haftmann@24219: in fold_consts the_const thms [] end; haftmann@24219: haftmann@26971: fun insts_of thy algebra tys sorts = haftmann@24219: let haftmann@24219: fun class_relation (x, _) _ = x; haftmann@24219: fun type_constructor tyco xs class = haftmann@26971: (tyco, class) :: (maps o maps) fst xs; haftmann@24219: fun type_variable (TVar (_, sort)) = map (pair []) sort haftmann@24219: | type_variable (TFree (_, sort)) = map (pair []) sort; haftmann@26971: fun of_sort_deriv ty sort = wenzelm@26939: Sorts.of_sort_derivation (Syntax.pp_global thy) algebra haftmann@24219: { class_relation = class_relation, type_constructor = type_constructor, haftmann@24219: type_variable = type_variable } haftmann@26517: (ty, sort) handle Sorts.CLASS_ERROR _ => [] (*permissive!*) haftmann@26971: in (flat o flat) (map2 of_sort_deriv tys sorts) end; haftmann@24219: haftmann@26971: fun meets_of thy algebra = haftmann@24219: let haftmann@26971: fun meet_of ty sort tab = haftmann@26971: Sorts.meet_sort algebra (ty, sort) tab haftmann@26971: handle Sorts.CLASS_ERROR _ => tab (*permissive!*); haftmann@26971: in fold2 meet_of end; haftmann@24219: haftmann@24219: haftmann@24219: (** graph algorithm **) haftmann@24219: haftmann@24219: val timing = ref false; haftmann@24219: haftmann@24219: local haftmann@24219: haftmann@26971: fun resort_thms thy algebra typ_of thms = haftmann@26971: let haftmann@26971: val cs = fold_consts (insert (op =)) thms []; haftmann@26971: fun meets (c, ty) = case typ_of c haftmann@26971: of SOME (vs, _) => haftmann@26971: meets_of thy algebra (Sign.const_typargs thy (c, ty)) (map snd vs) haftmann@26971: | NONE => I; haftmann@26971: val tab = fold meets cs Vartab.empty; haftmann@26971: in map (CodeUnit.inst_thm tab) thms end; haftmann@24219: haftmann@24219: fun resort_funcss thy algebra funcgr = haftmann@24219: let haftmann@24423: val typ_funcgr = try (fst o Graph.get_node funcgr); haftmann@26971: val resort_dep = apsnd (resort_thms thy algebra typ_funcgr); haftmann@26997: fun resort_rec typ_of (c, []) = (true, (c, [])) haftmann@26997: | resort_rec typ_of (c, thms as thm :: _) = if is_some (AxClass.inst_of_param thy c) haftmann@26997: then (true, (c, thms)) haftmann@26997: else let haftmann@26971: val (_, (vs, ty)) = CodeUnit.head_func thm; haftmann@26971: val thms' as thm' :: _ = resort_thms thy algebra typ_of thms haftmann@26971: val (_, (vs', ty')) = CodeUnit.head_func thm'; (*FIXME simplify check*) haftmann@26997: in (Sign.typ_equiv thy (ty, ty'), (c, thms')) end; haftmann@24219: fun resort_recs funcss = haftmann@24219: let haftmann@26971: fun typ_of c = case these (AList.lookup (op =) funcss c) haftmann@26971: of thm :: _ => (SOME o snd o CodeUnit.head_func) thm haftmann@26971: | [] => NONE; haftmann@26971: val (unchangeds, funcss') = split_list (map (resort_rec typ_of) funcss); haftmann@24219: val unchanged = fold (fn x => fn y => x andalso y) unchangeds true; haftmann@24219: in (unchanged, funcss') end; haftmann@24219: fun resort_rec_until funcss = haftmann@24219: let haftmann@24219: val (unchanged, funcss') = resort_recs funcss; haftmann@24219: in if unchanged then funcss' else resort_rec_until funcss' end; haftmann@24219: in map resort_dep #> resort_rec_until end; haftmann@24219: haftmann@24219: fun instances_of thy algebra insts = haftmann@24219: let haftmann@24219: val thy_classes = (#classes o Sorts.rep_algebra o Sign.classes_of) thy; haftmann@24835: fun all_classparams tyco class = wenzelm@24969: these (try (#params o AxClass.get_info thy) class) haftmann@26517: |> map_filter (fn (c, _) => try (AxClass.param_of_inst thy) (c, tyco)) haftmann@24219: in haftmann@24219: Symtab.empty haftmann@24219: |> fold (fn (tyco, class) => haftmann@24219: Symtab.map_default (tyco, []) (insert (op =) class)) insts haftmann@24835: |> (fn tab => Symtab.fold (fn (tyco, classes) => append (maps (all_classparams tyco) haftmann@24219: (Graph.all_succs thy_classes classes))) tab []) haftmann@24219: end; haftmann@24219: haftmann@24219: fun instances_of_consts thy algebra funcgr consts = haftmann@24219: let haftmann@26971: fun inst (cexpr as (c, ty)) = insts_of thy algebra haftmann@26971: (Sign.const_typargs thy (c, ty)) ((map snd o fst) (typ funcgr c)); haftmann@24219: in haftmann@24219: [] haftmann@24219: |> fold (fold (insert (op =)) o inst) consts haftmann@24219: |> instances_of thy algebra haftmann@24219: end; haftmann@24219: haftmann@24283: fun ensure_const' thy algebra funcgr const auxgr = haftmann@24423: if can (Graph.get_node funcgr) const haftmann@24219: then (NONE, auxgr) haftmann@24423: else if can (Graph.get_node auxgr) const haftmann@24219: then (SOME const, auxgr) haftmann@24219: else if is_some (Code.get_datatype_of_constr thy const) then haftmann@24219: auxgr haftmann@24423: |> Graph.new_node (const, []) haftmann@24219: |> pair (SOME const) haftmann@24219: else let haftmann@24219: val thms = Code.these_funcs thy const haftmann@24219: |> CodeUnit.norm_args haftmann@24219: |> CodeUnit.norm_varnames CodeName.purify_tvar CodeName.purify_var; haftmann@24219: val rhs = consts_of (const, thms); haftmann@24219: in haftmann@24219: auxgr haftmann@24423: |> Graph.new_node (const, thms) haftmann@24283: |> fold_map (ensure_const thy algebra funcgr) rhs haftmann@24423: |-> (fn rhs' => fold (fn SOME const' => Graph.add_edge (const, const') haftmann@24219: | NONE => I) rhs') haftmann@24219: |> pair (SOME const) haftmann@24219: end haftmann@24283: and ensure_const thy algebra funcgr const = haftmann@24219: let haftmann@24219: val timeap = if !timing haftmann@24219: then Output.timeap_msg ("time for " ^ CodeUnit.string_of_const thy const) haftmann@24219: else I; haftmann@24283: in timeap (ensure_const' thy algebra funcgr const) end; haftmann@24219: haftmann@24283: fun merge_funcss thy algebra raw_funcss funcgr = haftmann@24219: let haftmann@24219: val funcss = raw_funcss haftmann@24219: |> resort_funcss thy algebra funcgr haftmann@24423: |> filter_out (can (Graph.get_node funcgr) o fst); haftmann@24423: fun typ_func c [] = Code.default_typ thy c haftmann@26997: | typ_func c (thms as thm :: _) = (snd o CodeUnit.head_func) thm; haftmann@24219: fun add_funcs (const, thms) = haftmann@24423: Graph.new_node (const, (typ_func const thms, thms)); haftmann@24219: fun add_deps (funcs as (const, thms)) funcgr = haftmann@24219: let haftmann@24219: val deps = consts_of funcs; haftmann@24219: val insts = instances_of_consts thy algebra funcgr haftmann@24219: (fold_consts (insert (op =)) thms []); haftmann@24219: in haftmann@24219: funcgr haftmann@26997: |> ensure_consts thy algebra insts haftmann@24423: |> fold (curry Graph.add_edge const) deps haftmann@24423: |> fold (curry Graph.add_edge const) insts haftmann@24219: end; haftmann@24219: in haftmann@24219: funcgr haftmann@24219: |> fold add_funcs funcss haftmann@24219: |> fold add_deps funcss haftmann@24219: end haftmann@26997: and ensure_consts thy algebra cs funcgr = haftmann@24219: let haftmann@24423: val auxgr = Graph.empty haftmann@24283: |> fold (snd oo ensure_const thy algebra funcgr) cs; haftmann@24219: in haftmann@24219: funcgr haftmann@24283: |> fold (merge_funcss thy algebra) haftmann@24423: (map (AList.make (Graph.get_node auxgr)) haftmann@24423: (rev (Graph.strong_conn auxgr))) haftmann@26997: end; haftmann@24219: haftmann@24219: in haftmann@24219: haftmann@24219: (** retrieval interfaces **) haftmann@24219: haftmann@26997: val ensure_consts = ensure_consts; haftmann@24219: haftmann@26740: fun proto_eval thy cterm_of evaluator_fr evaluator proto_ct funcgr = haftmann@24219: let haftmann@26740: val ct = cterm_of proto_ct; wenzelm@26939: val _ = Sign.no_vars (Syntax.pp_global thy) (Thm.term_of ct); haftmann@24219: val _ = Term.fold_types (Type.no_tvars #> K I) (Thm.term_of ct) (); haftmann@26740: fun consts_of t = fold_aterms (fn Const c_ty => cons c_ty | _ => I) haftmann@26740: t []; haftmann@26740: val algebra = Code.coregular_algebra thy; haftmann@26740: val thm = Code.preprocess_conv ct; haftmann@26740: val ct' = Thm.rhs_of thm; haftmann@26740: val t' = Thm.term_of ct'; haftmann@26740: val consts = map fst (consts_of t'); haftmann@26740: val funcgr' = ensure_consts thy algebra consts funcgr; haftmann@26740: val (t'', evaluator') = apsnd evaluator_fr (evaluator t'); haftmann@26740: val consts' = consts_of t''; haftmann@26740: val dicts = instances_of_consts thy algebra funcgr' consts'; haftmann@26740: val funcgr'' = ensure_consts thy algebra dicts funcgr'; haftmann@26740: in (evaluator' thm funcgr'' t'', funcgr'') end; haftmann@26740: haftmann@26740: fun proto_eval_conv thy = haftmann@26740: let haftmann@26740: fun evaluator evaluator' thm1 funcgr t = haftmann@24219: let haftmann@26740: val thm2 = evaluator' funcgr t; haftmann@24219: val thm3 = Code.postprocess_conv (Thm.rhs_of thm2); haftmann@24219: in haftmann@26740: Thm.transitive thm1 (Thm.transitive thm2 thm3) handle THM _ => haftmann@26971: error ("could not construct evaluation proof:\n" wenzelm@26928: ^ (cat_lines o map Display.string_of_thm) [thm1, thm2, thm3]) haftmann@24219: end; haftmann@26740: in proto_eval thy I evaluator end; haftmann@24283: haftmann@26740: fun proto_eval_term thy = haftmann@24283: let haftmann@26740: fun evaluator evaluator' _ funcgr t = evaluator' funcgr t; haftmann@26740: in proto_eval thy (Thm.cterm_of thy) evaluator end; haftmann@24219: haftmann@24219: end; (*local*) haftmann@24219: haftmann@24219: structure Funcgr = CodeDataFun wenzelm@24713: ( haftmann@24219: type T = T; haftmann@24423: val empty = Graph.empty; haftmann@27609: fun purge _ cs funcgr = haftmann@27609: Graph.del_nodes ((Graph.all_preds funcgr haftmann@27609: o filter (can (Graph.get_node funcgr))) cs) funcgr; wenzelm@24713: ); haftmann@24219: haftmann@24219: fun make thy = haftmann@24423: Funcgr.change thy o ensure_consts thy (Code.coregular_algebra thy); haftmann@24219: haftmann@24219: fun eval_conv thy f = haftmann@26740: fst o Funcgr.change_yield thy o proto_eval_conv thy f; haftmann@24283: haftmann@24283: fun eval_term thy f = haftmann@26740: fst o Funcgr.change_yield thy o proto_eval_term thy f; haftmann@24219: haftmann@27103: haftmann@27103: (** diagnostic commands **) haftmann@27103: haftmann@27103: fun code_depgr thy [] = make thy [] haftmann@27103: | code_depgr thy consts = haftmann@27103: let haftmann@27103: val gr = make thy consts; haftmann@27103: val select = Graph.all_succs gr consts; haftmann@27103: in haftmann@27103: gr haftmann@27103: |> Graph.subgraph (member (op =) select) haftmann@27103: |> Graph.map_nodes ((apsnd o map) (AxClass.overload thy)) haftmann@27103: end; haftmann@27103: haftmann@27103: fun code_thms thy = Pretty.writeln o pretty thy o code_depgr thy; haftmann@27103: haftmann@27103: fun code_deps thy consts = haftmann@27103: let haftmann@27103: val gr = code_depgr thy consts; haftmann@27103: fun mk_entry (const, (_, (_, parents))) = haftmann@27103: let haftmann@27103: val name = CodeUnit.string_of_const thy const; haftmann@27103: val nameparents = map (CodeUnit.string_of_const thy) parents; haftmann@27103: in { name = name, ID = name, dir = "", unfold = true, haftmann@27103: path = "", parents = nameparents } haftmann@27103: end; haftmann@27103: val prgr = Graph.fold ((fn x => fn xs => xs @ [x]) o mk_entry) gr []; haftmann@27103: in Present.display_graph prgr end; haftmann@27103: haftmann@27103: local haftmann@27103: haftmann@27103: structure P = OuterParse haftmann@27103: and K = OuterKeyword haftmann@27103: haftmann@27103: fun code_thms_cmd thy = code_thms thy o op @ o CodeName.read_const_exprs thy; haftmann@27103: fun code_deps_cmd thy = code_deps thy o op @ o CodeName.read_const_exprs thy; haftmann@27103: haftmann@27103: in haftmann@27103: haftmann@27103: val _ = haftmann@27103: OuterSyntax.improper_command "code_thms" "print system of defining equations for code" OuterKeyword.diag haftmann@27103: (Scan.repeat P.term haftmann@27103: >> (fn cs => Toplevel.no_timing o Toplevel.unknown_theory haftmann@27103: o Toplevel.keep ((fn thy => code_thms_cmd thy cs) o Toplevel.theory_of))); haftmann@27103: haftmann@27103: val _ = haftmann@27103: OuterSyntax.improper_command "code_deps" "visualize dependencies of defining equations for code" OuterKeyword.diag haftmann@27103: (Scan.repeat P.term haftmann@27103: >> (fn cs => Toplevel.no_timing o Toplevel.unknown_theory haftmann@27103: o Toplevel.keep ((fn thy => code_deps_cmd thy cs) o Toplevel.theory_of))); haftmann@27103: haftmann@27103: end; haftmann@27103: haftmann@24219: end; (*struct*)