src/Pure/Tools/codegen_theorems.ML
changeset 20603 37dfd7af2746
parent 20602 96fa2cf465f5
child 20604 9dba9c7872c9
equal deleted inserted replaced
20602:96fa2cf465f5 20603:37dfd7af2746
     1 (*  Title:      Pure/Tools/codegen_theorems.ML
       
     2     ID:         $Id$
       
     3     Author:     Florian Haftmann, TU Muenchen
       
     4 
       
     5 Theorems used for code generation.
       
     6 *)
       
     7 
       
     8 signature CODEGEN_THEOREMS =
       
     9 sig
       
    10   val add_notify: ((string * typ) list option -> theory -> theory) -> theory -> theory;
       
    11   val add_preproc: (theory -> thm list -> thm list) -> theory -> theory;
       
    12   val add_fun_extr: (theory -> string * typ -> thm list) -> theory -> theory;
       
    13   val add_datatype_extr: (theory -> string
       
    14     -> (((string * sort) list * (string * typ list) list) * tactic) option)
       
    15     -> theory -> theory;
       
    16   val add_fun: thm -> theory -> theory;
       
    17   val del_fun: thm -> theory -> theory;
       
    18   val add_unfold: thm -> theory -> theory;
       
    19   val del_unfold: thm -> theory -> theory;
       
    20   val purge_defs: string * typ -> theory -> theory;
       
    21   val notify_dirty: theory -> theory;
       
    22 
       
    23   val extr_typ: theory -> thm -> typ;
       
    24   val rewrite_fun: thm list -> thm -> thm;
       
    25   val common_typ: theory -> (thm -> typ) -> thm list -> thm list;
       
    26   val preprocess: theory -> thm list -> thm list;
       
    27 
       
    28   val prove_freeness: theory -> tactic -> string
       
    29     -> (string * sort) list * (string * typ list) list -> thm list;
       
    30 
       
    31   type thmtab;
       
    32   val mk_thmtab: theory -> (string * typ) list -> thmtab;
       
    33   val get_dtyp_of_cons: thmtab -> string * typ -> string option;
       
    34   val get_dtyp_spec: thmtab -> string
       
    35     -> ((string * sort) list * (string * typ list) list) option;
       
    36   val get_fun_typs: thmtab -> ((string * typ list) * typ) list;
       
    37   val get_fun_thms: thmtab -> string * typ -> thm list;
       
    38 
       
    39   val pretty_funtab: theory -> thm list CodegenConsts.Consttab.table -> Pretty.T;
       
    40   val print_thms: theory -> unit;
       
    41 
       
    42   val init_obj: (thm * thm) * (thm * thm) -> theory -> theory;
       
    43   val debug: bool ref;
       
    44   val debug_msg: ('a -> string) -> 'a -> 'a;
       
    45 end;
       
    46 
       
    47 structure CodegenTheorems: CODEGEN_THEOREMS =
       
    48 struct
       
    49 
       
    50 (** preliminaries **)
       
    51 
       
    52 (* diagnostics *)
       
    53 
       
    54 val debug = ref false;
       
    55 fun debug_msg f x = (if !debug then Output.tracing (f x) else (); x);
       
    56 
       
    57 
       
    58 (* auxiliary *)
       
    59 
       
    60 fun getf_first [] _ = NONE
       
    61   | getf_first (f::fs) x = case f x
       
    62      of NONE => getf_first fs x
       
    63       | y as SOME x => y;
       
    64 
       
    65 fun getf_first_list [] x = []
       
    66   | getf_first_list (f::fs) x = case f x
       
    67      of [] => getf_first_list fs x
       
    68       | xs => xs;
       
    69 
       
    70 
       
    71 (* object logic setup *)
       
    72 
       
    73 structure CodegenTheoremsSetup = TheoryDataFun
       
    74 (struct
       
    75   val name = "Pure/codegen_theorems_setup";
       
    76   type T = ((string * thm) * ((string * string) * (string * string))) option;
       
    77   val empty = NONE;
       
    78   val copy = I;
       
    79   val extend = I;
       
    80   fun merge pp = merge_opt (eq_pair (eq_pair (op =) eq_thm)
       
    81     (eq_pair (eq_pair (op =) (op =)) (eq_pair (op =) (op =)))) : T * T -> T;
       
    82   fun print thy data = ();
       
    83 end);
       
    84 
       
    85 val _ = Context.add_setup CodegenTheoremsSetup.init;
       
    86 
       
    87 fun init_obj ((TrueI, FalseE), (conjI, atomize_eq)) thy =
       
    88   case CodegenTheoremsSetup.get thy
       
    89    of SOME _ => error "Code generator already set up for object logic"
       
    90     | NONE =>
       
    91         let
       
    92           fun strip_implies t = (Logic.strip_imp_prems t, Logic.strip_imp_concl t);
       
    93           fun dest_TrueI thm =
       
    94             Drule.plain_prop_of thm
       
    95             |> ObjectLogic.drop_judgment thy
       
    96             |> Term.dest_Const
       
    97             |> apsnd (
       
    98                   Term.dest_Type
       
    99                   #> fst
       
   100               );
       
   101           fun dest_FalseE thm =
       
   102             Drule.plain_prop_of thm
       
   103             |> Logic.dest_implies
       
   104             |> apsnd (
       
   105                  ObjectLogic.drop_judgment thy
       
   106                  #> Term.dest_Var
       
   107                )
       
   108             |> fst
       
   109             |> ObjectLogic.drop_judgment thy
       
   110             |> Term.dest_Const
       
   111             |> fst;
       
   112           fun dest_conjI thm =
       
   113             Drule.plain_prop_of thm
       
   114             |> strip_implies
       
   115             |> apfst (map (ObjectLogic.drop_judgment thy #> Term.dest_Var))
       
   116             |> apsnd (
       
   117                  ObjectLogic.drop_judgment thy
       
   118                  #> Term.strip_comb
       
   119                  #> apsnd (map Term.dest_Var)
       
   120                  #> apfst Term.dest_Const
       
   121                )
       
   122             |> (fn (v1, ((conj, _), v2)) => if v1 = v2 then conj else error "Wrong premise")
       
   123           fun dest_atomize_eq thm=
       
   124             Drule.plain_prop_of thm
       
   125             |> Logic.dest_equals
       
   126             |> apfst (
       
   127                  ObjectLogic.drop_judgment thy
       
   128                  #> Term.strip_comb
       
   129                  #> apsnd (map Term.dest_Var)
       
   130                  #> apfst Term.dest_Const
       
   131                )
       
   132             |> apsnd (
       
   133                  Logic.dest_equals
       
   134                  #> apfst Term.dest_Var
       
   135                  #> apsnd Term.dest_Var
       
   136                )
       
   137             |> (fn (((eq, _), v2), (v1a as (_, TVar (_, sort)), v1b)) =>
       
   138                  if [v1a, v1b] = v2 andalso sort = Sign.defaultS thy then eq else error "Wrong premise")
       
   139         in
       
   140           ((dest_TrueI TrueI, [dest_FalseE FalseE, dest_conjI conjI, dest_atomize_eq atomize_eq])
       
   141           handle _ => error "Bad code generator setup")
       
   142           |> (fn ((tr, b), [fl, con, eq]) => CodegenTheoremsSetup.put
       
   143                (SOME ((b, atomize_eq), ((tr, fl), (con, eq)))) thy)
       
   144         end;
       
   145 
       
   146 fun get_obj thy =
       
   147   case CodegenTheoremsSetup.get thy
       
   148    of SOME ((b, atomize), x) => ((Type (b, []), atomize) ,x)
       
   149     | NONE => error "No object logic setup for code theorems";
       
   150 
       
   151 fun mk_true thy =
       
   152   let
       
   153     val ((b, _), ((tr, fl), (con, eq))) = get_obj thy;
       
   154   in Const (tr, b) end;
       
   155 
       
   156 fun mk_false thy =
       
   157   let
       
   158     val ((b, _), ((tr, fl), (con, eq))) = get_obj thy;
       
   159   in Const (fl, b) end;
       
   160 
       
   161 fun mk_obj_conj thy (x, y) =
       
   162   let
       
   163     val ((b, _), ((tr, fl), (con, eq))) = get_obj thy;
       
   164   in Const (con, b --> b --> b) $ x $ y end;
       
   165 
       
   166 fun mk_obj_eq thy (x, y) =
       
   167   let
       
   168     val ((b, _), ((tr, fl), (con, eq))) = get_obj thy;
       
   169   in Const (eq, fastype_of x --> fastype_of y --> b) $ x $ y end;
       
   170 
       
   171 fun is_obj_eq thy c =
       
   172   let
       
   173     val ((b, _), ((tr, fl), (con, eq))) = get_obj thy;
       
   174   in c = eq end;
       
   175 
       
   176 fun mk_func thy ((x, y), rhs) =
       
   177   Logic.mk_equals (
       
   178     (mk_obj_eq thy (x, y)),
       
   179     rhs
       
   180   );
       
   181 
       
   182 
       
   183 (* theorem purification *)
       
   184 
       
   185 fun err_thm msg thm =
       
   186   error (msg ^ ": " ^ string_of_thm thm);
       
   187 
       
   188 val mk_rule =
       
   189   #mk o #mk_rews o snd o MetaSimplifier.rep_ss o Simplifier.simpset_of;
       
   190 
       
   191 fun abs_norm thy thm =
       
   192   let
       
   193     fun expvars t =
       
   194       let
       
   195         val lhs = (fst o Logic.dest_equals) t;
       
   196         val tys = (fst o strip_type o fastype_of) lhs;
       
   197         val used = fold_aterms (fn Var ((v, _), _) => insert (op =) v | _ => I) lhs [];
       
   198         val vs = Name.invent_list used "x" (length tys);
       
   199       in
       
   200         map2 (fn v => fn ty => Var ((v, 0), ty)) vs tys
       
   201       end;
       
   202     fun expand ct thm =
       
   203       Thm.combination thm (Thm.reflexive ct);
       
   204     fun beta_norm thm =
       
   205       thm
       
   206       |> prop_of
       
   207       |> Logic.dest_equals
       
   208       |> fst
       
   209       |> cterm_of thy
       
   210       |> Thm.beta_conversion true
       
   211       |> Thm.symmetric
       
   212       |> (fn thm' => Thm.transitive thm' thm);
       
   213   in
       
   214     thm
       
   215     |> fold (expand o cterm_of thy) ((expvars o prop_of) thm)
       
   216     |> beta_norm
       
   217   end;
       
   218 
       
   219 fun canonical_tvars thy thm =
       
   220   let
       
   221     fun mk_inst (v_i as (v, i), (v', sort)) (s as (maxidx, set, acc)) =
       
   222       if v = v' orelse member (op =) set v then s
       
   223         else let
       
   224           val ty = TVar (v_i, sort)
       
   225         in
       
   226           (maxidx + 1, v :: set,
       
   227             (ctyp_of thy ty, ctyp_of thy (TVar ((v', maxidx), sort))) :: acc)
       
   228         end;
       
   229     fun tvars_of thm = (fold_types o fold_atyps)
       
   230       (fn TVar (v_i as (v, i), sort) => cons (v_i, (CodegenNames.purify_var v, sort))
       
   231         | _ => I) (prop_of thm) [];
       
   232     val maxidx = Thm.maxidx_of thm + 1;
       
   233     val (_, _, inst) = fold mk_inst (tvars_of thm) (maxidx + 1, [], []);
       
   234   in Thm.instantiate (inst, []) thm end;
       
   235 
       
   236 fun canonical_vars thy thm =
       
   237   let
       
   238     fun mk_inst (v_i as (v, i), (v', ty)) (s as (maxidx, set, acc)) =
       
   239       if v = v' orelse member (op =) set v then s
       
   240         else let
       
   241           val t = if i = ~1 then Free (v, ty) else Var (v_i, ty)
       
   242         in
       
   243           (maxidx + 1,  v :: set,
       
   244             (cterm_of thy t, cterm_of thy (Var ((v', maxidx), ty))) :: acc)
       
   245         end;
       
   246     fun vars_of thm = fold_aterms
       
   247       (fn Var (v_i as (v, i), ty) => cons (v_i, (CodegenNames.purify_var v, ty))
       
   248         | _ => I) (prop_of thm) [];
       
   249     val maxidx = Thm.maxidx_of thm + 1;
       
   250     val (_, _, inst) = fold mk_inst (vars_of thm) (maxidx + 1, [], []);
       
   251   in Thm.instantiate ([], inst) thm end;
       
   252 
       
   253 fun drop_redundant thy eqs =
       
   254   let
       
   255     val matches = curry (Pattern.matches thy o
       
   256       pairself (fst o Logic.dest_equals o prop_of))
       
   257     fun drop eqs [] = eqs
       
   258       | drop eqs (eq::eqs') =
       
   259           drop (eq::eqs) (filter_out (matches eq) eqs')
       
   260   in drop [] eqs end;
       
   261 
       
   262 fun drop_refl thy = filter_out (is_equal o Term.fast_term_ord o Logic.dest_equals
       
   263   o ObjectLogic.drop_judgment thy o Drule.plain_prop_of);
       
   264 
       
   265 fun make_eq thy =
       
   266   let
       
   267     val ((_, atomize), _) = get_obj thy;
       
   268   in rewrite_rule [atomize] end;
       
   269 
       
   270 fun dest_eq thy thm =
       
   271   case try (make_eq thy #> Drule.plain_prop_of
       
   272    #> ObjectLogic.drop_judgment thy #> Logic.dest_equals) thm
       
   273    of SOME eq => (eq, thm)
       
   274     | NONE => err_thm "Not an equation" thm;
       
   275 
       
   276 fun dest_fun thy thm =
       
   277   let
       
   278     fun dest_fun' ((lhs, _), thm) =
       
   279       case try (dest_Const o fst o strip_comb) lhs
       
   280        of SOME (c, ty) => (c, (ty, thm))
       
   281         | NONE => err_thm "Not a function equation" thm;
       
   282   in
       
   283     thm
       
   284     |> dest_eq thy
       
   285     |> dest_fun'
       
   286   end;
       
   287 
       
   288 
       
   289 
       
   290 (** theory data **)
       
   291 
       
   292 (* data structures *)
       
   293 
       
   294 structure Consttab = CodegenConsts.Consttab;
       
   295 
       
   296 fun merge' eq (xys as (xs, ys)) =
       
   297   if eq_list eq (xs, ys) then (false, xs) else (true, merge eq xys);
       
   298 
       
   299 fun alist_merge' eq_key eq (xys as (xs, ys)) =
       
   300   if eq_list (eq_pair eq_key eq) (xs, ys) then (false, xs) else (true, AList.merge eq_key eq xys);
       
   301 
       
   302 fun list_consttab_join' eq (xyt as (xt, yt)) =
       
   303   let
       
   304     val xc = Consttab.keys xt;
       
   305     val yc = Consttab.keys yt;
       
   306     val zc = filter (member CodegenConsts.eq_const yc) xc;
       
   307     val wc = subtract (op =) zc xc @ subtract (op =) zc yc;
       
   308     fun same_thms c = if eq_list eq_thm ((the o Consttab.lookup xt) c, (the o Consttab.lookup yt) c)
       
   309       then NONE else SOME c;
       
   310   in (wc @ map_filter same_thms zc, Consttab.join (K (merge eq)) xyt) end;
       
   311 
       
   312 datatype notify = Notify of (serial * ((string * typ) list option -> theory -> theory)) list;
       
   313 
       
   314 val mk_notify = Notify;
       
   315 fun map_notify f (Notify notify) = mk_notify (f notify);
       
   316 fun merge_notify pp (Notify notify1, Notify notify2) =
       
   317   mk_notify (AList.merge (op =) (K true) (notify1, notify2));
       
   318 
       
   319 datatype preproc = Preproc of {
       
   320   preprocs: (serial * (theory -> thm list -> thm list)) list,
       
   321   unfolds: thm list
       
   322 };
       
   323 
       
   324 fun mk_preproc (preprocs, unfolds) =
       
   325   Preproc { preprocs = preprocs, unfolds = unfolds };
       
   326 fun map_preproc f (Preproc { preprocs, unfolds }) =
       
   327   mk_preproc (f (preprocs, unfolds));
       
   328 fun merge_preproc _ (Preproc { preprocs = preprocs1, unfolds = unfolds1 },
       
   329   Preproc { preprocs = preprocs2, unfolds = unfolds2 }) =
       
   330     let
       
   331       val (dirty1, preprocs) = alist_merge' (op =) (K true) (preprocs1, preprocs2);
       
   332       val (dirty2, unfolds) = merge' eq_thm (unfolds1, unfolds2);
       
   333     in (dirty1 orelse dirty2, mk_preproc (preprocs, unfolds)) end;
       
   334 
       
   335 datatype extrs = Extrs of {
       
   336   funs: (serial * (theory -> string * typ -> thm list)) list,
       
   337   datatypes: (serial * (theory -> string -> (((string * sort) list * (string * typ list) list) * tactic) option)) list
       
   338 };
       
   339 
       
   340 fun mk_extrs (funs, datatypes) =
       
   341   Extrs { funs = funs, datatypes = datatypes };
       
   342 fun map_extrs f (Extrs { funs, datatypes }) =
       
   343   mk_extrs (f (funs, datatypes));
       
   344 fun merge_extrs _ (Extrs { funs = funs1, datatypes = datatypes1 },
       
   345   Extrs { funs = funs2, datatypes = datatypes2 }) =
       
   346     let
       
   347       val (dirty1, funs) = alist_merge' (op =) (K true) (funs1, funs2);
       
   348       val (dirty2, datatypes) = alist_merge' (op =) (K true) (datatypes1, datatypes2);
       
   349     in (dirty1 orelse dirty2, mk_extrs (funs, datatypes)) end;
       
   350 
       
   351 datatype funthms = Funthms of {
       
   352   dirty: string list,
       
   353   funs: thm list Consttab.table
       
   354 };
       
   355 
       
   356 fun mk_funthms (dirty, funs) =
       
   357   Funthms { dirty = dirty, funs = funs };
       
   358 fun map_funthms f (Funthms { dirty, funs }) =
       
   359   mk_funthms (f (dirty, funs));
       
   360 fun merge_funthms _ (Funthms { dirty = dirty1, funs = funs1 },
       
   361   Funthms { dirty = dirty2, funs = funs2 }) =
       
   362     let
       
   363       val (dirty3, funs) = list_consttab_join' eq_thm (funs1, funs2);
       
   364     in mk_funthms (merge (op =) (merge (op =) (dirty1, dirty2), map fst dirty3), funs) end;
       
   365 
       
   366 datatype T = T of {
       
   367   dirty: bool,
       
   368   notify: notify,
       
   369   preproc: preproc,
       
   370   extrs: extrs,
       
   371   funthms: funthms
       
   372 };
       
   373 
       
   374 fun mk_T ((dirty, notify), (preproc, (extrs, funthms))) =
       
   375   T { dirty = dirty, notify = notify, preproc= preproc, extrs = extrs, funthms = funthms };
       
   376 fun map_T f (T { dirty, notify, preproc, extrs, funthms }) =
       
   377   mk_T (f ((dirty, notify), (preproc, (extrs, funthms))));
       
   378 fun merge_T pp (T { dirty = dirty1, notify = notify1, preproc = preproc1, extrs = extrs1, funthms = funthms1 },
       
   379   T { dirty = dirty2, notify = notify2, preproc = preproc2, extrs = extrs2, funthms = funthms2 }) =
       
   380     let
       
   381       val (dirty3, preproc) = merge_preproc pp (preproc1, preproc2);
       
   382       val (dirty4, extrs) = merge_extrs pp (extrs1, extrs2);
       
   383     in
       
   384       mk_T ((dirty1 orelse dirty2 orelse dirty3 orelse dirty4, merge_notify pp (notify1, notify2)),
       
   385         (preproc, (extrs, merge_funthms pp (funthms1, funthms2))))
       
   386     end;
       
   387 
       
   388 
       
   389 (* setup *)
       
   390 
       
   391 structure CodegenTheoremsData = TheoryDataFun
       
   392 (struct
       
   393   val name = "Pure/codegen_theorems_data";
       
   394   type T = T;
       
   395   val empty = mk_T ((false, mk_notify []), (mk_preproc ([], []),
       
   396     (mk_extrs ([], []), mk_funthms ([], Consttab.empty))));
       
   397   val copy = I;
       
   398   val extend = I;
       
   399   val merge = merge_T;
       
   400   fun print (thy : theory) (data : T) =
       
   401     let
       
   402       val pretty_thm = ProofContext.pretty_thm (ProofContext.init thy);
       
   403       val funthms = (fn T { funthms, ... } => funthms) data;
       
   404       val funs = (Consttab.dest o (fn Funthms { funs, ... } => funs)) funthms;
       
   405       val preproc = (fn T { preproc, ... } => preproc) data;
       
   406       val unfolds = (fn Preproc { unfolds, ... } => unfolds) preproc;
       
   407     in
       
   408       (Pretty.writeln o Pretty.block o Pretty.fbreaks) ([
       
   409         Pretty.str "code generation theorems:",
       
   410         Pretty.str "function theorems:" ] @
       
   411           map (fn (c, thms) =>
       
   412             (Pretty.block o Pretty.fbreaks) (
       
   413               (Pretty.str o CodegenConsts.string_of_const thy) c  :: map pretty_thm (rev thms)
       
   414             )
       
   415           ) funs @ [
       
   416         Pretty.block (
       
   417           Pretty.str "inlined theorems:"
       
   418           :: Pretty.fbrk
       
   419           :: (Pretty.fbreaks o map pretty_thm) unfolds
       
   420       )])
       
   421     end;
       
   422 end);
       
   423 
       
   424 val _ = Context.add_setup CodegenTheoremsData.init;
       
   425 val print_thms = CodegenTheoremsData.print;
       
   426 
       
   427 
       
   428 (* accessors *)
       
   429 
       
   430 local
       
   431   val the_preproc = (fn T { preproc = Preproc preproc, ... } => preproc) o CodegenTheoremsData.get;
       
   432   val the_extrs = (fn T { extrs = Extrs extrs, ... } => extrs) o CodegenTheoremsData.get;
       
   433   val the_funthms = (fn T { funthms = Funthms funthms, ... } => funthms) o CodegenTheoremsData.get;
       
   434 in
       
   435   val is_dirty = (fn T { dirty = dirty, ... } => dirty) o CodegenTheoremsData.get;
       
   436   val the_dirty_consts = (fn { dirty = dirty, ... } => dirty) o the_funthms;
       
   437   val the_notify = (fn T { notify = Notify notify, ... } => map snd notify) o CodegenTheoremsData.get;
       
   438   val the_preprocs = (fn { preprocs, ... } => map snd preprocs) o the_preproc;
       
   439   val the_unfolds = (fn { unfolds, ... } => unfolds) o the_preproc;
       
   440   val the_funs_extrs = (fn { funs, ... } => map snd funs) o the_extrs;
       
   441   val the_datatypes_extrs = (fn { datatypes, ... } => map snd datatypes) o the_extrs;
       
   442   val the_funs = (fn { funs, ... } => funs) o the_funthms;
       
   443 end (*local*);
       
   444 
       
   445 val map_data = CodegenTheoremsData.map o map_T;
       
   446 
       
   447 (* notifiers *)
       
   448 
       
   449 fun all_typs thy c =
       
   450   let
       
   451     val c_tys = (map (pair c o #lhs o snd) o Defs.specifications_of (Theory.defs_of thy)) c;
       
   452   in (c, Sign.the_const_type thy c) :: map (CodegenConsts.typ_of_typinst thy) c_tys end;
       
   453 
       
   454 fun add_notify f =
       
   455   map_data (fn ((dirty, notify), x) =>
       
   456     ((dirty, notify |> map_notify (cons (serial (), f))), x));
       
   457 
       
   458 fun get_reset_dirty thy =
       
   459   let
       
   460     val dirty = is_dirty thy;
       
   461     val dirty_const = if dirty then [] else the_dirty_consts thy;
       
   462   in
       
   463     thy
       
   464     |> map_data (fn ((_, notify), (procs, (extrs, funthms))) =>
       
   465          ((false, notify), (procs, (extrs, funthms |> map_funthms (fn (_, funs) => ([], funs))))))
       
   466     |> pair (dirty, dirty_const)
       
   467   end;
       
   468 
       
   469 fun notify_all c thy =
       
   470   thy
       
   471   |> get_reset_dirty
       
   472   |-> (fn (true, _) => fold (fn f => f NONE) (the_notify thy)
       
   473         | (false, cs) => let val cs' = case c of NONE => cs | SOME c => insert (op =) c cs
       
   474             in fold (fn f => f (SOME (maps (all_typs thy) cs'))) (the_notify thy) end);
       
   475 
       
   476 fun notify_dirty thy =
       
   477   thy
       
   478   |> get_reset_dirty
       
   479   |-> (fn (true, _) => fold (fn f => f NONE) (the_notify thy)
       
   480         | (false, cs) => fold (fn f => f (SOME (maps (all_typs thy) cs))) (the_notify thy));
       
   481 
       
   482 
       
   483 (* modifiers *)
       
   484 
       
   485 fun add_preproc f =
       
   486   map_data (fn (x, (preproc, y)) =>
       
   487     (x, (preproc |> map_preproc (fn (preprocs, unfolds) => ((serial (), f) :: preprocs, unfolds)), y)))
       
   488   #> notify_all NONE;
       
   489 
       
   490 fun add_fun_extr f =
       
   491   map_data (fn (x, (preproc, (extrs, funthms))) =>
       
   492     (x, (preproc, (extrs |> map_extrs (fn (funs, datatypes) =>
       
   493       ((serial (), f) :: funs, datatypes)), funthms))))
       
   494   #> notify_all NONE;
       
   495 
       
   496 fun add_datatype_extr f =
       
   497   map_data (fn (x, (preproc, (extrs, funthms))) =>
       
   498     (x, (preproc, (extrs |> map_extrs (fn (funs, datatypes) =>
       
   499       (funs, (serial (), f) :: datatypes)), funthms))))
       
   500   #> notify_all NONE;
       
   501 
       
   502 fun add_fun thm thy =
       
   503   case dest_fun thy thm
       
   504    of (c, (ty, _)) =>
       
   505     thy
       
   506     |> map_data (fn (x, (preproc, (extrs, funthms))) =>
       
   507         (x, (preproc, (extrs, funthms |> map_funthms (fn (dirty, funs) =>
       
   508           (dirty, funs |> Consttab.map_default (CodegenConsts.norminst_of_typ thy (c, ty), []) (cons thm)))))))
       
   509     |> notify_all (SOME c);
       
   510 
       
   511 fun del_fun thm thy =
       
   512   case dest_fun thy thm
       
   513    of (c, (ty, _)) =>
       
   514     thy
       
   515     |> map_data (fn (x, (preproc, (extrs, funthms))) =>
       
   516         (x, (preproc, (extrs, funthms |> map_funthms (fn (dirty, funs) =>
       
   517           (dirty, funs |> Consttab.map_entry (CodegenConsts.norminst_of_typ thy (c, ty)) (remove eq_thm thm)))))))
       
   518     |> notify_all (SOME c);
       
   519 
       
   520 fun add_unfold thm thy =
       
   521   thy
       
   522   |> tap (fn thy => dest_eq thy thm)
       
   523   |> map_data (fn (x, (preproc, y)) =>
       
   524        (x, (preproc |> map_preproc (fn (preprocs, unfolds) =>
       
   525          (preprocs, thm :: unfolds)), y)))
       
   526   |> notify_all NONE;
       
   527 
       
   528 fun del_unfold thm =
       
   529   map_data (fn (x, (preproc, y)) =>
       
   530        (x, (preproc |> map_preproc (fn (preprocs, unfolds) =>
       
   531          (preprocs, remove eq_thm thm unfolds)), y)))
       
   532   #> notify_all NONE;
       
   533 
       
   534 fun purge_defs (c, ty) thy =
       
   535   thy
       
   536   |> map_data (fn (x, (preproc, (extrs, funthms))) =>
       
   537       (x, (preproc, (extrs, funthms |> map_funthms (fn (dirty, funs) =>
       
   538         (dirty, funs |> Consttab.update (CodegenConsts.norminst_of_typ thy (c, ty), [])))))))
       
   539   |> notify_all (SOME c);
       
   540 
       
   541 
       
   542 
       
   543 (** theorem handling **)
       
   544 
       
   545 (* preprocessing *)
       
   546 
       
   547 fun extr_typ thy thm = case dest_fun thy thm
       
   548  of (_, (ty, _)) => ty;
       
   549 
       
   550 fun rewrite_fun rewrites thm =
       
   551   let
       
   552     val rewrite = Tactic.rewrite true rewrites;
       
   553     val (ct_eq, [ct_lhs, ct_rhs]) = (Drule.strip_comb o cprop_of) thm;
       
   554     val Const ("==", _) = term_of ct_eq;
       
   555     val (ct_f, ct_args) = Drule.strip_comb ct_lhs;
       
   556     val rhs' = rewrite ct_rhs;
       
   557     val args' = map rewrite ct_args;
       
   558     val lhs' = Thm.symmetric (fold (fn th1 => fn th2 => Thm.combination th2 th1)
       
   559       args' (Thm.reflexive ct_f));
       
   560   in
       
   561     Thm.transitive (Thm.transitive lhs' thm) rhs'
       
   562   end handle Bind => raise ERROR "rewrite_fun"
       
   563 
       
   564 fun common_typ thy _ [] = []
       
   565   | common_typ thy _ [thm] = [thm]
       
   566   | common_typ thy extract_typ thms =
       
   567       let
       
   568         fun incr_thm thm max =
       
   569           let
       
   570             val thm' = incr_indexes max thm;
       
   571             val max' = (maxidx_of_typ o fastype_of o Drule.plain_prop_of) thm' + 1;
       
   572           in (thm', max') end;
       
   573         val (thms', maxidx) = fold_map incr_thm thms 0;
       
   574         val (ty1::tys) = map extract_typ thms;
       
   575         fun unify ty env = Sign.typ_unify thy (ty1, ty) env
       
   576           handle Type.TUNIFY =>
       
   577             error ("Type unificaton failed, while unifying function equations\n"
       
   578             ^ (cat_lines o map Display.string_of_thm) thms
       
   579             ^ "\nwith types\n"
       
   580             ^ (cat_lines o map (Sign.string_of_typ thy)) (ty1 :: tys));
       
   581         val (env, _) = fold unify tys (Vartab.empty, maxidx)
       
   582         val instT = Vartab.fold (fn (x_i, (sort, ty)) =>
       
   583           cons (Thm.ctyp_of thy (TVar (x_i, sort)), Thm.ctyp_of thy ty)) env [];
       
   584       in map (Thm.instantiate (instT, [])) thms end;
       
   585 
       
   586 fun preprocess thy thms =
       
   587   let
       
   588     fun burrow_thms f [] = []
       
   589       | burrow_thms f thms =
       
   590           thms
       
   591           |> Conjunction.intr_list
       
   592           |> f
       
   593           |> Conjunction.elim_list;
       
   594     fun cmp_thms (thm1, thm2) =
       
   595       not (Sign.typ_instance thy (extr_typ thy thm1, extr_typ thy thm2));
       
   596     fun unvarify thms =
       
   597       #2 (#1 (Variable.import true thms (ProofContext.init thy)));
       
   598     val unfold_thms = map (make_eq thy) (the_unfolds thy);
       
   599   in
       
   600     thms
       
   601     |> map (make_eq thy)
       
   602     |> map (Thm.transfer thy)
       
   603     |> fold (fn f => f thy) (the_preprocs thy)
       
   604     |> map (rewrite_fun unfold_thms)
       
   605     |> debug_msg (fn _ => "[cg_thm] sorting")
       
   606     |> debug_msg (commas o map string_of_thm)
       
   607     |> sort (make_ord cmp_thms)
       
   608     |> debug_msg (fn _ => "[cg_thm] common_typ")
       
   609     |> debug_msg (commas o map string_of_thm)
       
   610     |> common_typ thy (extr_typ thy)
       
   611     |> debug_msg (fn _ => "[cg_thm] abs_norm")
       
   612     |> debug_msg (commas o map string_of_thm)
       
   613     |> map (abs_norm thy)
       
   614     |> drop_refl thy
       
   615     |> burrow_thms (
       
   616         debug_msg (fn _ => "[cg_thm] canonical tvars")
       
   617         #> debug_msg (string_of_thm)
       
   618         #> canonical_tvars thy
       
   619         #> debug_msg (fn _ => "[cg_thm] canonical vars")
       
   620         #> debug_msg (string_of_thm)
       
   621         #> canonical_vars thy
       
   622         #> debug_msg (fn _ => "[cg_thm] zero indices")
       
   623         #> debug_msg (string_of_thm)
       
   624         #> Drule.zero_var_indexes
       
   625        )
       
   626     |> drop_redundant thy
       
   627     |> debug_msg (fn _ => "[cg_thm] preprocessing done")
       
   628   end;
       
   629 
       
   630 
       
   631 (* retrieval *)
       
   632 
       
   633 fun get_funs thy (c, ty) =
       
   634   let
       
   635     val _ = debug_msg (fn _ => "[cg_thm] const (1) " ^ c ^ " :: " ^ Sign.string_of_typ thy ty) ()
       
   636     val postprocess_typ = case AxClass.class_of_param thy c
       
   637      of NONE => map_filter (fn (_, (ty', thm)) =>
       
   638           if Sign.typ_instance thy (ty, ty')
       
   639           then SOME thm else debug_msg (fn _ => "[cg_thm] dropping " ^ string_of_thm thm) NONE)
       
   640       | SOME _ => let
       
   641           (*FIXME make this more elegant*)
       
   642           val ty' = CodegenConsts.typ_of_classop thy (CodegenConsts.norminst_of_typ thy (c, ty));
       
   643           val ct = Thm.cterm_of thy (Const (c, ty'));
       
   644           val thm' = Thm.reflexive ct;
       
   645         in map (snd o snd) #> cons thm' #> common_typ thy (extr_typ thy) #> tl end;
       
   646     fun get_funs (c, ty) =
       
   647       (these o Consttab.lookup (the_funs thy) o CodegenConsts.norminst_of_typ thy) (c, ty)
       
   648       |> debug_msg (fn _ => "[cg_thm] trying funs")
       
   649       |> map (dest_fun thy)
       
   650       |> postprocess_typ;
       
   651     fun get_extr (c, ty) =
       
   652       getf_first_list (map (fn f => f thy) (the_funs_extrs thy)) (c, ty)
       
   653       |> debug_msg (fn _ => "[cg_thm] trying extr")
       
   654       |> map (dest_fun thy)
       
   655       |> postprocess_typ;
       
   656     fun get_spec (c, ty) =
       
   657       (CodegenConsts.find_def thy o CodegenConsts.norminst_of_typ thy) (c, ty)
       
   658       |> debug_msg (fn _ => "[cg_thm] trying spec")
       
   659       |> Option.mapPartial (fn ((_, name), _) => try (Thm.get_axiom_i thy) name)
       
   660       |> the_list
       
   661       |> map_filter (try (dest_fun thy))
       
   662       |> postprocess_typ;
       
   663   in
       
   664     getf_first_list [get_funs, get_extr, get_spec] (c, ty)
       
   665     |> debug_msg (fn _ => "[cg_thm] const (2) " ^ c ^ " :: " ^ Sign.string_of_typ thy ty)
       
   666     |> preprocess thy
       
   667   end;
       
   668 
       
   669 fun prove_freeness thy tac dtco vs_cos =
       
   670   let
       
   671     val truh = mk_true thy;
       
   672     val fals = mk_false thy;
       
   673     fun mk_lhs vs ((co1, tys1), (co2, tys2)) =
       
   674       let
       
   675         val dty = Type (dtco, map TFree vs);
       
   676         val (xs1, xs2) = chop (length tys1) (Name.invent_list [] "a" (length tys1 + length tys2));
       
   677         val frees1 = map2 (fn x => fn ty => Free (x, ty)) xs1 tys1;
       
   678         val frees2 = map2 (fn x => fn ty => Free (x, ty)) xs2 tys2;
       
   679         fun zip_co co xs tys = list_comb (Const (co,
       
   680           tys ---> dty), map2 (fn x => fn ty => Free (x, ty)) xs tys);
       
   681       in
       
   682         ((frees1, frees2), (zip_co co1 xs1 tys1, zip_co co2 xs2 tys2))
       
   683       end;
       
   684     fun mk_rhs [] [] = truh
       
   685       | mk_rhs xs ys = foldr1 (mk_obj_conj thy) (map2 (curry (mk_obj_eq thy)) xs ys);
       
   686     fun mk_eq vs (args as ((co1, _), (co2, _))) (inj, dist) =
       
   687       if co1 = co2
       
   688         then let
       
   689           val ((fs1, fs2), lhs) = mk_lhs vs args;
       
   690           val rhs = mk_rhs fs1 fs2;
       
   691         in (mk_func thy (lhs, rhs) :: inj, dist) end
       
   692         else let
       
   693           val (_, lhs) = mk_lhs vs args;
       
   694         in (inj, mk_func thy (lhs, fals) :: dist) end;
       
   695     fun mk_eqs (vs, cos) =
       
   696       let val cos' = rev cos
       
   697       in (op @) (fold (mk_eq vs) (product cos' cos') ([], [])) end;
       
   698     val ts = (map (ObjectLogic.ensure_propT thy) o mk_eqs) vs_cos;
       
   699     fun prove t = if !quick_and_dirty then SkipProof.make_thm thy (Logic.varify t)
       
   700       else Goal.prove_global thy [] [] t (K tac);
       
   701   in map prove ts end;
       
   702 
       
   703 fun get_datatypes thy dtco =
       
   704   let
       
   705     val _ = debug_msg (fn _ => "[cg_thm] datatype " ^ dtco) ()
       
   706   in
       
   707     case getf_first (map (fn f => f thy) (the_datatypes_extrs thy)) dtco
       
   708      of NONE => NONE
       
   709       | SOME (vs_cos, tac) => SOME (vs_cos, prove_freeness thy tac dtco vs_cos)
       
   710   end;
       
   711 
       
   712 fun get_eq thy (c, ty) =
       
   713   if is_obj_eq thy c
       
   714   then case strip_type ty
       
   715    of (Type (tyco, _) :: _, _) =>
       
   716      (case get_datatypes thy tyco
       
   717        of SOME (_, thms) => thms
       
   718         | _ => [])
       
   719     | _ => []
       
   720   else [];
       
   721 
       
   722 fun check_thms c thms =
       
   723   let
       
   724     fun check_head_lhs thm (lhs, rhs) =
       
   725       case strip_comb lhs
       
   726        of (Const (c', _), _) => if c' = c then ()
       
   727            else error ("Illegal function equation for " ^ quote c
       
   728              ^ ", actually defining " ^ quote c' ^ ": " ^ Display.string_of_thm thm)
       
   729         | _ => error ("Illegal function equation: " ^ Display.string_of_thm thm);
       
   730     fun check_vars_lhs thm (lhs, rhs) =
       
   731       if has_duplicates (op =)
       
   732           (fold_aterms (fn Free (v, _) => cons v | _ => I) lhs [])
       
   733       then error ("Repeated variables on left hand side of function equation:"
       
   734         ^ Display.string_of_thm thm)
       
   735       else ();
       
   736     fun check_vars_rhs thm (lhs, rhs) =
       
   737       if null (subtract (op =)
       
   738         (fold_aterms (fn Free (v, _) => cons v | _ => I) lhs [])
       
   739         (fold_aterms (fn Free (v, _) => cons v | _ => I) rhs []))
       
   740       then ()
       
   741       else error ("Free variables on right hand side of function equation:"
       
   742         ^ Display.string_of_thm thm)
       
   743     val tts = map (Logic.dest_equals o Logic.unvarify o Thm.prop_of) thms;
       
   744   in
       
   745     (map2 check_head_lhs thms tts; map2 check_vars_lhs thms tts;
       
   746       map2 check_vars_rhs thms tts; thms)
       
   747   end;
       
   748 
       
   749 structure Consttab = CodegenConsts.Consttab;
       
   750 type thmtab = (theory * (thm list Consttab.table
       
   751   * string Consttab.table)
       
   752   * ((string * sort) list * (string * typ list) list) Symtab.table);
       
   753 
       
   754 fun thmtab_empty thy = (thy, (Consttab.empty, Consttab.empty),
       
   755   Symtab.empty);
       
   756 
       
   757 fun get_dtyp_of_cons (thy, (_, dtcotab), _) =
       
   758   Consttab.lookup dtcotab o CodegenConsts.norminst_of_typ thy;
       
   759 
       
   760 fun get_dtyp_spec (_, _, dttab) =
       
   761   Symtab.lookup dttab;
       
   762 
       
   763 fun has_fun_thms (thy, (funtab, _), _) =
       
   764   is_some o Consttab.lookup funtab o CodegenConsts.norminst_of_typ thy;
       
   765 
       
   766 fun get_fun_thms (thy, (funtab, _), _) (c_ty as (c, _)) =
       
   767   (check_thms c o these o Consttab.lookup funtab
       
   768     o CodegenConsts.norminst_of_typ thy) c_ty;
       
   769 
       
   770 fun get_fun_typs (thy, (funtab, dtcotab), _) =
       
   771   (Consttab.dest funtab
       
   772   |> map (fn (c, thm :: _) => (c, extr_typ thy thm)
       
   773            | (c as (name, _), []) => (c, case AxClass.class_of_param thy name
       
   774          of SOME class => (case ClassPackage.the_consts_sign thy class of (v, cs) =>
       
   775               (Logic.varifyT o map_type_tfree (fn u as (w, _) =>
       
   776                 if w = v then TFree (v, [class]) else TFree u))
       
   777               ((the o AList.lookup (op =) cs) name))
       
   778           | NONE => Sign.the_const_type thy name)))
       
   779   @ (Consttab.keys dtcotab
       
   780   |> AList.make (Sign.const_instance thy));
       
   781 
       
   782 fun pretty_funtab thy funtab =
       
   783   funtab
       
   784   |> CodegenConsts.Consttab.dest
       
   785   |> map (fn (c, thms) =>
       
   786        (Pretty.block o Pretty.fbreaks) (
       
   787          (Pretty.str o CodegenConsts.string_of_const thy) c
       
   788          :: map Display.pretty_thm thms
       
   789        ))
       
   790   |> Pretty.chunks;
       
   791 
       
   792 fun constrain_funtab thy funtab =
       
   793   let
       
   794     fun max k [] = k
       
   795       | max k (l::ls) = max (if k < l then l else k) ls;
       
   796     fun mk_consttyps funtab =
       
   797       CodegenConsts.Consttab.empty
       
   798       |> CodegenConsts.Consttab.fold (fn (c, thm :: _) =>
       
   799            CodegenConsts.Consttab.update_new (c, extr_typ thy thm) | (_, []) => I) funtab
       
   800     fun mk_typescheme_of typtab (c, ty) =
       
   801       CodegenConsts.Consttab.lookup typtab (CodegenConsts.norminst_of_typ thy (c, ty));
       
   802     fun incr_indices (c, thms) maxidx =
       
   803       let
       
   804         val thms' = map (Thm.incr_indexes maxidx) thms;
       
   805         val maxidx' = Int.max
       
   806           (maxidx, max ~1 (map Thm.maxidx_of thms') + 1);
       
   807       in (thms', maxidx') end;
       
   808     fun consts_of_eqs thms =
       
   809       let
       
   810         fun terms_of_eq thm =
       
   811           let
       
   812             val (lhs, rhs) = (Logic.dest_equals o Drule.plain_prop_of) thm
       
   813           in rhs :: (snd o strip_comb) lhs end;
       
   814       in (fold o fold_aterms) (fn Const c => insert (eq_pair (op =) (Type.eq_type Vartab.empty)) c | _ => I)
       
   815         (maps terms_of_eq thms) []
       
   816       end;
       
   817     val typscheme_of =
       
   818       mk_typescheme_of (mk_consttyps funtab);
       
   819     val tsig = Sign.tsig_of thy;
       
   820     fun unify_const (c, ty) (env, maxidx) =
       
   821       case typscheme_of (c, ty)
       
   822        of SOME ty_decl => let
       
   823             (*val _ = writeln "UNIFY";
       
   824             val _ = writeln (CodegenConsts.string_of_const_typ thy (c, ty))*)
       
   825             val ty_decl' = Logic.incr_tvar maxidx ty_decl;
       
   826             (*val _ = writeln "WITH";
       
   827             val _ = writeln (CodegenConsts.string_of_const_typ thy (c, ty_decl'))*)
       
   828             val maxidx' = Int.max (Term.maxidx_of_typ ty_decl' + 1, maxidx);
       
   829             (*val _ = writeln ("  " ^ string_of_int maxidx ^ " +> " ^ string_of_int maxidx');*)
       
   830           in Type.unify tsig (ty_decl', ty) (env, maxidx') end
       
   831         | NONE => (env, maxidx);
       
   832     fun apply_unifier unif [] = []
       
   833       | apply_unifier unif (thms as thm :: _) =
       
   834           let
       
   835             val ty = extr_typ thy thm;
       
   836             val ty' = Envir.norm_type unif ty;
       
   837             val env = Type.typ_match (Sign.tsig_of thy) (ty, ty') Vartab.empty;
       
   838             val inst = Thm.instantiate (Vartab.fold (fn (x_i, (sort, ty)) =>
       
   839               cons (Thm.ctyp_of thy (TVar (x_i, sort)), Thm.ctyp_of thy ty)) env [], []);
       
   840           in map (Drule.zero_var_indexes o inst) thms end;
       
   841 (*     val _ = writeln "(1)";  *)
       
   842 (*     val _ = (Pretty.writeln o pretty_funtab thy) funtab;  *)
       
   843     val (funtab', maxidx) =
       
   844       CodegenConsts.Consttab.fold_map incr_indices funtab 0;
       
   845 (*     val _ = writeln "(2)";
       
   846  *     val _ = (Pretty.writeln o pretty_funtab thy) funtab';
       
   847  *)
       
   848     val (unif, _) =
       
   849       CodegenConsts.Consttab.fold (fold unify_const o consts_of_eqs o snd)
       
   850         funtab' (Vartab.empty, maxidx);
       
   851 (*     val _ = writeln "(3)";  *)
       
   852     val funtab'' =
       
   853       CodegenConsts.Consttab.map (apply_unifier unif) funtab';
       
   854 (*     val _ = writeln "(4)";
       
   855  *     val _ = (Pretty.writeln o pretty_funtab thy) funtab'';
       
   856  *)
       
   857   in funtab'' end;
       
   858 
       
   859 fun mk_thmtab thy cs =
       
   860   let
       
   861     fun add_tycos (Type (tyco, tys)) = insert (op =) tyco #> fold add_tycos tys
       
   862       | add_tycos _ = I;
       
   863     fun consts_of ts =
       
   864       Consttab.empty
       
   865       |> (fold o fold_aterms)
       
   866            (fn Const c_ty => Consttab.update (CodegenConsts.norminst_of_typ thy c_ty, ())
       
   867              | _ => I) ts
       
   868       |> Consttab.keys;
       
   869     fun add_dtyps_of_type ty thmtab =
       
   870       let
       
   871         val tycos = add_tycos ty [];
       
   872         val tycos_new = filter (is_none o get_dtyp_spec thmtab) tycos;
       
   873         fun add_dtyp_spec dtco (dtyp_spec as (vs, cs)) ((thy, (funtab, dtcotab), dttab)) =
       
   874           let
       
   875             fun mk_co (c, tys) =
       
   876               CodegenConsts.norminst_of_typ thy (c, Logic.varifyT (tys ---> Type (dtco, map TFree vs)));
       
   877           in
       
   878             (thy, (funtab, dtcotab |> fold (fn c_tys =>
       
   879               Consttab.update_new (mk_co c_tys, dtco)) cs),
       
   880               dttab |> Symtab.update_new (dtco, dtyp_spec))
       
   881           end;
       
   882       in
       
   883         thmtab
       
   884         |> fold (fn tyco => case get_datatypes thy tyco
       
   885              of SOME (dtyp_spec, _) => add_dtyp_spec tyco dtyp_spec
       
   886               | NONE => I) tycos_new
       
   887       end;
       
   888     fun known thmtab (c, ty) =
       
   889       is_some (get_dtyp_of_cons thmtab (c, ty)) orelse has_fun_thms thmtab (c, ty);
       
   890     fun add_funthms (c, ty) (thmtab as (thy, (funtab, dtcotab), algebra_dttab))=
       
   891       if known thmtab (c, ty) then thmtab
       
   892       else let
       
   893         val thms = get_funs thy (c, ty)
       
   894         val cs_dep = (consts_of o map Thm.prop_of) thms;
       
   895       in
       
   896         (thy, (funtab |> Consttab.update_new (CodegenConsts.norminst_of_typ thy (c, ty), thms)
       
   897         , dtcotab), algebra_dttab)
       
   898         |> fold add_c cs_dep
       
   899       end
       
   900     and add_c (c_tys as (c, tys)) thmtab =
       
   901       thmtab
       
   902       |> add_dtyps_of_type (snd (CodegenConsts.typ_of_typinst thy c_tys))
       
   903       |> fold (add_funthms o CodegenConsts.typ_of_typinst thy)
       
   904            (CodegenConsts.insts_of_classop thy c_tys);
       
   905   in
       
   906     thmtab_empty thy
       
   907     |> fold (add_c o CodegenConsts.norminst_of_typ thy) cs
       
   908     |> (fn (a, (funtab, b), c) => (a, (funtab |> constrain_funtab thy, b), c))
       
   909   end;
       
   910 
       
   911 
       
   912 
       
   913 (** code attributes and setup **)
       
   914 
       
   915 local
       
   916   fun add_simple_attribute (name, f) =
       
   917     (Codegen.add_attribute name o (Scan.succeed o Thm.declaration_attribute))
       
   918       (Context.map_theory o f);
       
   919 in
       
   920   val _ = map (Context.add_setup o add_simple_attribute) [
       
   921     ("func", add_fun),
       
   922     ("nofun", del_fun),
       
   923     ("unfold", (fn thm => Codegen.add_unfold thm #> add_unfold thm)),
       
   924     ("inline", add_unfold),
       
   925     ("noinline", del_unfold)
       
   926   ]
       
   927 end; (*local*)
       
   928 
       
   929 val _ = Context.add_setup (add_fun_extr get_eq);
       
   930 
       
   931 end; (*struct*)