| 20600 |      1 | (*  Title:      Pure/Tools/codegen_data.ML
 | 
|  |      2 |     ID:         $Id$
 | 
|  |      3 |     Author:     Florian Haftmann, TU Muenchen
 | 
|  |      4 | 
 | 
| 20855 |      5 | Abstract executable content of theory. Management of data dependent on
 | 
|  |      6 | executabl content.
 | 
| 20600 |      7 | *)
 | 
|  |      8 | 
 | 
|  |      9 | signature CODEGEN_DATA =
 | 
|  |     10 | sig
 | 
|  |     11 |   type lthms = thm list Susp.T;
 | 
|  |     12 |   val lazy: (unit -> thm list) -> lthms
 | 
|  |     13 |   val eval_always: bool ref
 | 
|  |     14 | 
 | 
|  |     15 |   val add_func: thm -> theory -> theory
 | 
|  |     16 |   val del_func: thm -> theory -> theory
 | 
|  |     17 |   val add_funcl: CodegenConsts.const * lthms -> theory -> theory
 | 
|  |     18 |   val add_datatype: string * (((string * sort) list * (string * typ list) list) * lthms)
 | 
|  |     19 |     -> theory -> theory
 | 
|  |     20 |   val del_datatype: string -> theory -> theory
 | 
|  |     21 |   val add_inline: thm -> theory -> theory
 | 
|  |     22 |   val del_inline: thm -> theory -> theory
 | 
| 20844 |     23 |   val add_inline_proc: (theory -> cterm list -> thm list) -> theory -> theory
 | 
|  |     24 |   val add_constrains: (theory -> term list -> (indexname * sort) list) -> theory -> theory
 | 
| 20600 |     25 |   val add_preproc: (theory -> thm list -> thm list) -> theory -> theory
 | 
|  |     26 |   val these_funcs: theory -> CodegenConsts.const -> thm list
 | 
|  |     27 |   val get_datatype: theory -> string
 | 
|  |     28 |     -> ((string * sort) list * (string * typ list) list) option
 | 
|  |     29 |   val get_datatype_of_constr: theory -> CodegenConsts.const -> string option
 | 
|  |     30 | 
 | 
|  |     31 |   val print_thms: theory -> unit
 | 
|  |     32 | 
 | 
|  |     33 |   val typ_func: theory -> thm -> typ
 | 
| 20937 |     34 |   val typ_funcs: theory -> CodegenConsts.const * thm list -> typ
 | 
| 20600 |     35 |   val rewrite_func: thm list -> thm -> thm
 | 
| 20844 |     36 |   val preprocess_cterm: theory -> (string * typ -> typ) -> cterm -> thm * cterm
 | 
| 20600 |     37 | 
 | 
| 20844 |     38 |   val trace: bool ref
 | 
| 20600 |     39 |   val strict_functyp: bool ref
 | 
|  |     40 | end;
 | 
|  |     41 | 
 | 
|  |     42 | signature PRIVATE_CODEGEN_DATA =
 | 
|  |     43 | sig
 | 
|  |     44 |   include CODEGEN_DATA
 | 
|  |     45 |   type data
 | 
|  |     46 |   structure CodeData: THEORY_DATA
 | 
|  |     47 |   val declare: string -> Object.T -> (Pretty.pp -> Object.T * Object.T -> Object.T)
 | 
| 20937 |     48 |     -> (theory option -> CodegenConsts.const list option -> Object.T -> Object.T) -> serial
 | 
| 20600 |     49 |   val init: serial -> theory -> theory
 | 
|  |     50 |   val get: serial -> (Object.T -> 'a) -> data -> 'a
 | 
|  |     51 |   val put: serial -> ('a -> Object.T) -> 'a -> data -> data
 | 
|  |     52 | end;
 | 
|  |     53 | 
 | 
|  |     54 | structure CodegenData : PRIVATE_CODEGEN_DATA =
 | 
|  |     55 | struct
 | 
|  |     56 | 
 | 
|  |     57 | (** diagnostics **)
 | 
|  |     58 | 
 | 
| 20844 |     59 | val trace = ref false;
 | 
|  |     60 | fun tracing f x = (if !trace then Output.tracing (f x) else (); x);
 | 
| 20600 |     61 | 
 | 
|  |     62 | 
 | 
|  |     63 | 
 | 
|  |     64 | (** lazy theorems, certificate theorems **)
 | 
|  |     65 | 
 | 
|  |     66 | type lthms = thm list Susp.T;
 | 
|  |     67 | val eval_always = ref false;
 | 
|  |     68 | 
 | 
|  |     69 | fun lazy f = if !eval_always
 | 
|  |     70 |   then Susp.value (f ())
 | 
|  |     71 |   else Susp.delay f;
 | 
|  |     72 | 
 | 
|  |     73 | fun string_of_lthms r = case Susp.peek r
 | 
|  |     74 |  of SOME thms => (map string_of_thm o rev) thms
 | 
|  |     75 |   | NONE => ["[...]"];
 | 
|  |     76 | 
 | 
|  |     77 | fun pretty_lthms ctxt r = case Susp.peek r
 | 
|  |     78 |  of SOME thms => (map (ProofContext.pretty_thm ctxt) o rev) thms
 | 
|  |     79 |   | NONE => [Pretty.str "[...]"];
 | 
|  |     80 | 
 | 
| 20844 |     81 | fun certificate thy f r =
 | 
| 20600 |     82 |   case Susp.peek r
 | 
| 20844 |     83 |    of SOME thms => (Susp.value o f thy) thms
 | 
|  |     84 |      | NONE => let
 | 
|  |     85 |           val thy_ref = Theory.self_ref thy;
 | 
|  |     86 |         in lazy (fn () => (f (Theory.deref thy_ref) o Susp.force) r) end;
 | 
| 20600 |     87 | 
 | 
|  |     88 | fun merge' _ ([], []) = (false, [])
 | 
|  |     89 |   | merge' _ ([], ys) = (true, ys)
 | 
|  |     90 |   | merge' eq (xs, ys) = fold_rev
 | 
|  |     91 |       (fn y => fn (t, xs) => (t orelse not (member eq xs y), insert eq y xs)) ys (false, xs);
 | 
|  |     92 | 
 | 
|  |     93 | fun merge_alist eq_key eq (xys as (xs, ys)) =
 | 
|  |     94 |   if eq_list (eq_pair eq_key eq) (xs, ys)
 | 
|  |     95 |   then (false, xs)
 | 
|  |     96 |   else (true, AList.merge eq_key eq xys);
 | 
|  |     97 | 
 | 
|  |     98 | val merge_thms = merge' eq_thm;
 | 
|  |     99 | 
 | 
|  |    100 | fun merge_lthms (r1, r2) =
 | 
|  |    101 |   if Susp.same (r1, r2) then (false, r1)
 | 
|  |    102 |   else case Susp.peek r1
 | 
|  |    103 |    of SOME [] => (true, r2)
 | 
|  |    104 |     | _ => case Susp.peek r2
 | 
|  |    105 |        of SOME [] => (true, r1)
 | 
|  |    106 |         | _ => (apsnd (lazy o K)) (merge_thms (Susp.force r1, Susp.force r2));
 | 
|  |    107 | 
 | 
|  |    108 | 
 | 
|  |    109 | 
 | 
|  |    110 | (** code theorems **)
 | 
|  |    111 | 
 | 
| 20844 |    112 | (* making rewrite theorems *)
 | 
| 20600 |    113 | 
 | 
|  |    114 | fun bad_thm msg thm =
 | 
|  |    115 |   error (msg ^ ": " ^ string_of_thm thm);
 | 
|  |    116 | 
 | 
| 20844 |    117 | fun check_rew thy thm =
 | 
|  |    118 |   let
 | 
|  |    119 |     val (lhs, rhs) = (Logic.dest_equals o Thm.prop_of) thm;
 | 
|  |    120 |     fun vars_of t = fold_aterms
 | 
|  |    121 |      (fn Var (v, _) => insert (op =) v
 | 
|  |    122 |        | Free _ => bad_thm "Illegal free variable in rewrite theorem" thm
 | 
|  |    123 |        | _ => I) t [];
 | 
|  |    124 |     fun tvars_of t = fold_term_types
 | 
|  |    125 |      (fn _ => fold_atyps (fn TVar (v, _) => insert (op =) v
 | 
|  |    126 |                           | TFree _ => bad_thm "Illegal free type variable in rewrite theorem" thm)) t [];
 | 
|  |    127 |     val lhs_vs = vars_of lhs;
 | 
|  |    128 |     val rhs_vs = vars_of rhs;
 | 
|  |    129 |     val lhs_tvs = tvars_of lhs;
 | 
|  |    130 |     val rhs_tvs = tvars_of lhs;
 | 
|  |    131 |     val _ = if null (subtract (op =) lhs_vs rhs_vs)
 | 
|  |    132 |       then ()
 | 
|  |    133 |       else bad_thm "Free variables on right hand side of rewrite theorems" thm
 | 
|  |    134 |     val _ = if null (subtract (op =) lhs_tvs rhs_tvs)
 | 
|  |    135 |       then ()
 | 
|  |    136 |       else bad_thm "Free type variables on right hand side of rewrite theorems" thm
 | 
|  |    137 |   in thm end;
 | 
|  |    138 | 
 | 
|  |    139 | fun mk_rew thy thm =
 | 
|  |    140 |   let
 | 
|  |    141 |     val thms = (#mk o #mk_rews o snd o MetaSimplifier.rep_ss o Simplifier.simpset_of) thy thm;
 | 
|  |    142 |   in
 | 
|  |    143 |     map (check_rew thy) thms
 | 
|  |    144 |   end;
 | 
|  |    145 | 
 | 
|  |    146 | 
 | 
|  |    147 | (* making function theorems *)
 | 
|  |    148 | 
 | 
| 20600 |    149 | fun typ_func thy = snd o dest_Const o fst o strip_comb o fst 
 | 
|  |    150 |   o Logic.dest_equals o ObjectLogic.drop_judgment thy o Drule.plain_prop_of;
 | 
|  |    151 | 
 | 
| 20844 |    152 | val strict_functyp = ref true;
 | 
|  |    153 | 
 | 
|  |    154 | fun dest_func thy = apfst dest_Const o strip_comb o Envir.beta_eta_contract
 | 
|  |    155 |   o fst o Logic.dest_equals o ObjectLogic.drop_judgment thy o Drule.plain_prop_of;
 | 
|  |    156 | 
 | 
|  |    157 | fun mk_head thy thm =
 | 
|  |    158 |   ((CodegenConsts.norm_of_typ thy o fst o dest_func thy) thm, thm);
 | 
| 20600 |    159 | 
 | 
| 20844 |    160 | fun check_func verbose thy thm = case try (dest_func thy) thm
 | 
|  |    161 |  of SOME (c_ty as (c, ty), args) =>
 | 
|  |    162 |       let
 | 
|  |    163 |         val _ =
 | 
|  |    164 |           if has_duplicates (op =)
 | 
|  |    165 |             ((fold o fold_aterms) (fn Var (v, _) => cons v | _ => I) args [])
 | 
|  |    166 |           then bad_thm "Repeated variables on left hand side of function equation" thm
 | 
|  |    167 |           else ()
 | 
|  |    168 |         val is_classop = (is_some o AxClass.class_of_param thy) c;
 | 
|  |    169 |         val const = CodegenConsts.norm_of_typ thy c_ty;
 | 
|  |    170 |         val ty_decl = CodegenConsts.disc_typ_of_const thy
 | 
|  |    171 |           (snd o CodegenConsts.typ_of_inst thy) const;
 | 
|  |    172 |         val string_of_typ = setmp show_sorts true (Sign.string_of_typ thy);
 | 
|  |    173 |       in if Sign.typ_equiv thy (ty_decl, ty)
 | 
|  |    174 |         then (const, thm)
 | 
|  |    175 |         else (if is_classop orelse (!strict_functyp andalso not
 | 
|  |    176 |           (Sign.typ_equiv thy (Type.strip_sorts ty_decl, Type.strip_sorts ty)))
 | 
|  |    177 |           then error else (if verbose then warning else K ()) #> K (const, thm))
 | 
|  |    178 |           ("Type\n" ^ string_of_typ ty
 | 
|  |    179 |            ^ "\nof function theorem\n"
 | 
|  |    180 |            ^ string_of_thm thm
 | 
|  |    181 |            ^ "\nis strictly less general than declared function type\n"
 | 
|  |    182 |            ^ string_of_typ ty_decl) 
 | 
|  |    183 |       end
 | 
|  |    184 |   | NONE => bad_thm "Not a function equation" thm;
 | 
|  |    185 | 
 | 
|  |    186 | fun check_typ_classop thy thm =
 | 
|  |    187 |   let
 | 
|  |    188 |     val (c_ty as (c, ty), _) = dest_func thy thm;  
 | 
|  |    189 |   in case AxClass.class_of_param thy c
 | 
|  |    190 |    of SOME class => let
 | 
|  |    191 |         val const = CodegenConsts.norm_of_typ thy c_ty;
 | 
|  |    192 |         val ty_decl = CodegenConsts.disc_typ_of_const thy
 | 
|  |    193 |             (snd o CodegenConsts.typ_of_inst thy) const;
 | 
|  |    194 |         val string_of_typ = setmp show_sorts true (Sign.string_of_typ thy);
 | 
|  |    195 |       in if Sign.typ_equiv thy (ty_decl, ty)
 | 
|  |    196 |         then thm
 | 
|  |    197 |         else error
 | 
|  |    198 |           ("Type\n" ^ string_of_typ ty
 | 
|  |    199 |            ^ "\nof function theorem\n"
 | 
|  |    200 |            ^ string_of_thm thm
 | 
|  |    201 |            ^ "\nis strictly less general than declared function type\n"
 | 
|  |    202 |            ^ string_of_typ ty_decl)
 | 
|  |    203 |       end
 | 
|  |    204 |     | NONE => thm
 | 
|  |    205 |   end;
 | 
| 20600 |    206 | 
 | 
|  |    207 | fun mk_func thy raw_thm =
 | 
| 20844 |    208 |   mk_rew thy raw_thm
 | 
|  |    209 |   |> map (check_func true thy);
 | 
| 20600 |    210 | 
 | 
|  |    211 | fun get_prim_def_funcs thy c =
 | 
|  |    212 |   let
 | 
|  |    213 |     fun constrain thm0 thm = case AxClass.class_of_param thy (fst c)
 | 
|  |    214 |      of SOME _ =>
 | 
|  |    215 |           let
 | 
| 20704 |    216 |             val ty_decl = CodegenConsts.disc_typ_of_classop thy c;
 | 
| 20600 |    217 |             val max = maxidx_of_typ ty_decl + 1;
 | 
|  |    218 |             val thm = Thm.incr_indexes max thm;
 | 
|  |    219 |             val ty = typ_func thy thm;
 | 
|  |    220 |             val (env, _) = Sign.typ_unify thy (ty_decl, ty) (Vartab.empty, max);
 | 
|  |    221 |             val instT = Vartab.fold (fn (x_i, (sort, ty)) =>
 | 
|  |    222 |               cons (Thm.ctyp_of thy (TVar (x_i, sort)), Thm.ctyp_of thy ty)) env [];
 | 
|  |    223 |           in Thm.instantiate (instT, []) thm end
 | 
|  |    224 |       | NONE => thm
 | 
|  |    225 |   in case CodegenConsts.find_def thy c
 | 
|  |    226 |    of SOME ((_, thm), _) =>
 | 
|  |    227 |         thm
 | 
|  |    228 |         |> Thm.transfer thy
 | 
|  |    229 |         |> try (map snd o mk_func thy)
 | 
|  |    230 |         |> these
 | 
|  |    231 |         |> map (constrain thm)
 | 
|  |    232 |     | NONE => []
 | 
|  |    233 |   end;
 | 
|  |    234 | 
 | 
|  |    235 | 
 | 
|  |    236 | (* pairs of (selected, deleted) function theorems *)
 | 
|  |    237 | 
 | 
|  |    238 | type sdthms = lthms * thm list;
 | 
|  |    239 | 
 | 
|  |    240 | fun add_drop_redundant thm thms =
 | 
|  |    241 |   let
 | 
|  |    242 |     val thy = Context.check_thy (Thm.theory_of_thm thm);
 | 
|  |    243 |     val pattern = (fst o Logic.dest_equals o Drule.plain_prop_of) thm;
 | 
|  |    244 |     fun matches thm' = if (curry (Pattern.matches thy) pattern o
 | 
|  |    245 |       fst o Logic.dest_equals o Drule.plain_prop_of) thm'
 | 
|  |    246 |       then (warning ("Dropping redundant function theorem\n" ^ string_of_thm thm'); true)
 | 
|  |    247 |       else false
 | 
|  |    248 |   in thm :: filter_out matches thms end;
 | 
|  |    249 | 
 | 
|  |    250 | fun add_thm thm (sels, dels) =
 | 
|  |    251 |   (Susp.value (add_drop_redundant thm (Susp.force sels)), remove eq_thm thm dels);
 | 
|  |    252 | 
 | 
|  |    253 | fun add_lthms lthms (sels, []) =
 | 
|  |    254 |       (lazy (fn () => fold add_drop_redundant
 | 
|  |    255 |           (Susp.force lthms) (Susp.force sels)), [])
 | 
|  |    256 |   | add_lthms lthms (sels, dels) =
 | 
|  |    257 |       fold add_thm (Susp.force lthms) (sels, dels);
 | 
|  |    258 | 
 | 
|  |    259 | fun del_thm thm (sels, dels) =
 | 
|  |    260 |   (Susp.value (remove eq_thm thm (Susp.force sels)), thm :: dels);
 | 
|  |    261 | 
 | 
|  |    262 | fun pretty_sdthms ctxt (sels, _) = pretty_lthms ctxt sels;
 | 
|  |    263 | 
 | 
|  |    264 | fun merge_sdthms c ((sels1, dels1), (sels2, dels2)) =
 | 
|  |    265 |   let
 | 
|  |    266 |     val (dels_t, dels) = merge_thms (dels1, dels2);
 | 
|  |    267 |   in if dels_t
 | 
|  |    268 |     then let
 | 
|  |    269 |       val (_, sels) = merge_thms (Susp.force sels1, subtract eq_thm dels1 (Susp.force sels2))
 | 
|  |    270 |       val (_, dels) = merge_thms (dels1, subtract eq_thm (Susp.force sels1) dels2)
 | 
|  |    271 |     in (true, ((lazy o K) sels, dels)) end
 | 
|  |    272 |     else let
 | 
|  |    273 |       val (sels_t, sels) = merge_lthms (sels1, sels2)
 | 
|  |    274 |     in (sels_t, (sels, dels)) end
 | 
|  |    275 |   end;
 | 
|  |    276 | 
 | 
|  |    277 | 
 | 
|  |    278 | (** data structures **)
 | 
|  |    279 | 
 | 
|  |    280 | structure Consttab = CodegenConsts.Consttab;
 | 
|  |    281 | 
 | 
|  |    282 | datatype preproc = Preproc of {
 | 
|  |    283 |   inlines: thm list,
 | 
| 20844 |    284 |   inline_procs: (serial * (theory -> cterm list -> thm list)) list,
 | 
|  |    285 |   constrains: (serial * (theory -> term list -> (indexname * sort) list)) list,
 | 
| 20600 |    286 |   preprocs: (serial * (theory -> thm list -> thm list)) list
 | 
|  |    287 | };
 | 
|  |    288 | 
 | 
| 20844 |    289 | fun mk_preproc ((inlines, inline_procs), (constrains, preprocs)) =
 | 
|  |    290 |   Preproc { inlines = inlines, inline_procs = inline_procs, constrains = constrains, preprocs = preprocs };
 | 
|  |    291 | fun map_preproc f (Preproc { inlines, inline_procs, constrains, preprocs }) =
 | 
|  |    292 |   mk_preproc (f ((inlines, inline_procs), (constrains, preprocs)));
 | 
|  |    293 | fun merge_preproc (Preproc { inlines = inlines1, inline_procs = inline_procs1, constrains = constrains1 , preprocs = preprocs1 },
 | 
|  |    294 |   Preproc { inlines = inlines2, inline_procs = inline_procs2, constrains = constrains2 , preprocs = preprocs2 }) =
 | 
| 20600 |    295 |     let
 | 
|  |    296 |       val (touched1, inlines) = merge_thms (inlines1, inlines2);
 | 
| 20844 |    297 |       val (touched2, inline_procs) = merge_alist (op =) (K true) (inline_procs1, inline_procs2);
 | 
|  |    298 |       val (touched3, constrains) = merge_alist (op =) (K true) (constrains1, constrains2);
 | 
|  |    299 |       val (touched4, preprocs) = merge_alist (op =) (K true) (preprocs1, preprocs2);
 | 
|  |    300 |     in (touched1 orelse touched2 orelse touched3 orelse touched4,
 | 
|  |    301 |       mk_preproc ((inlines, inline_procs), (constrains, preprocs))) end;
 | 
| 20600 |    302 | 
 | 
|  |    303 | fun join_func_thms (tabs as (tab1, tab2)) =
 | 
|  |    304 |   let
 | 
|  |    305 |     val cs1 = Consttab.keys tab1;
 | 
|  |    306 |     val cs2 = Consttab.keys tab2;
 | 
|  |    307 |     val cs' = filter (member CodegenConsts.eq_const cs2) cs1;
 | 
|  |    308 |     val cs'' = subtract (op =) cs' cs1 @ subtract (op =) cs' cs2;
 | 
|  |    309 |     val cs''' = ref [] : CodegenConsts.const list ref;
 | 
|  |    310 |     fun merge c x = let val (touched, thms') = merge_sdthms c x in
 | 
|  |    311 |       (if touched then cs''' := cons c (!cs''') else (); thms') end;
 | 
|  |    312 |   in (cs'' @ !cs''', Consttab.join merge tabs) end;
 | 
|  |    313 | fun merge_funcs (thms1, thms2) =
 | 
|  |    314 |     let
 | 
|  |    315 |       val (consts, thms) = join_func_thms (thms1, thms2);
 | 
|  |    316 |     in (SOME consts, thms) end;
 | 
|  |    317 | 
 | 
|  |    318 | val eq_string = op = : string * string -> bool;
 | 
|  |    319 | fun eq_dtyp (((vs1, cs1), _), ((vs2, cs2), _)) = 
 | 
|  |    320 |   gen_eq_set (eq_pair eq_string (gen_eq_set eq_string)) (vs1, vs2)
 | 
|  |    321 |     andalso gen_eq_set (eq_pair eq_string (eq_list (is_equal o Term.typ_ord))) (cs1, cs2);
 | 
|  |    322 | fun merge_dtyps (tabs as (tab1, tab2)) =
 | 
|  |    323 |   let
 | 
|  |    324 |     val tycos1 = Symtab.keys tab1;
 | 
|  |    325 |     val tycos2 = Symtab.keys tab2;
 | 
|  |    326 |     val tycos' = filter (member eq_string tycos2) tycos1;
 | 
| 20844 |    327 |     val touched = not (gen_eq_set (op =) (tycos1, tycos2) andalso
 | 
|  |    328 |       gen_eq_set (eq_pair (op =) (eq_dtyp))
 | 
| 20600 |    329 |       (AList.make (the o Symtab.lookup tab1) tycos',
 | 
| 20844 |    330 |        AList.make (the o Symtab.lookup tab2) tycos'));
 | 
| 20600 |    331 |   in (touched, Symtab.merge (K true) tabs) end;
 | 
|  |    332 | 
 | 
|  |    333 | datatype spec = Spec of {
 | 
|  |    334 |   funcs: sdthms Consttab.table,
 | 
|  |    335 |   dconstrs: string Consttab.table,
 | 
|  |    336 |   dtyps: (((string * sort) list * (string * typ list) list) * lthms) Symtab.table
 | 
|  |    337 | };
 | 
|  |    338 | 
 | 
|  |    339 | fun mk_spec ((funcs, dconstrs), dtyps) =
 | 
|  |    340 |   Spec { funcs = funcs, dconstrs = dconstrs, dtyps = dtyps };
 | 
|  |    341 | fun map_spec f (Spec { funcs = funcs, dconstrs = dconstrs, dtyps = dtyps }) =
 | 
|  |    342 |   mk_spec (f ((funcs, dconstrs), dtyps));
 | 
|  |    343 | fun merge_spec (Spec { funcs = funcs1, dconstrs = dconstrs1, dtyps = dtyps1 },
 | 
|  |    344 |   Spec { funcs = funcs2, dconstrs = dconstrs2, dtyps = dtyps2 }) =
 | 
|  |    345 |   let
 | 
|  |    346 |     val (touched_cs, funcs) = merge_funcs (funcs1, funcs2);
 | 
|  |    347 |     val dconstrs = Consttab.merge (K true) (dconstrs1, dconstrs2);
 | 
|  |    348 |     val (touched', dtyps) = merge_dtyps (dtyps1, dtyps2);
 | 
|  |    349 |     val touched = if touched' then NONE else touched_cs;
 | 
|  |    350 |   in (touched, mk_spec ((funcs, dconstrs), dtyps)) end;
 | 
|  |    351 | 
 | 
|  |    352 | datatype exec = Exec of {
 | 
|  |    353 |   preproc: preproc,
 | 
|  |    354 |   spec: spec
 | 
|  |    355 | };
 | 
|  |    356 | 
 | 
|  |    357 | fun mk_exec (preproc, spec) =
 | 
|  |    358 |   Exec { preproc = preproc, spec = spec };
 | 
|  |    359 | fun map_exec f (Exec { preproc = preproc, spec = spec }) =
 | 
|  |    360 |   mk_exec (f (preproc, spec));
 | 
|  |    361 | fun merge_exec (Exec { preproc = preproc1, spec = spec1 },
 | 
|  |    362 |   Exec { preproc = preproc2, spec = spec2 }) =
 | 
|  |    363 |   let
 | 
|  |    364 |     val (touched', preproc) = merge_preproc (preproc1, preproc2);
 | 
|  |    365 |     val (touched_cs, spec) = merge_spec (spec1, spec2);
 | 
|  |    366 |     val touched = if touched' then NONE else touched_cs;
 | 
|  |    367 |   in (touched, mk_exec (preproc, spec)) end;
 | 
| 20844 |    368 | val empty_exec = mk_exec (mk_preproc (([], []), ([], [])),
 | 
| 20600 |    369 |   mk_spec ((Consttab.empty, Consttab.empty), Symtab.empty));
 | 
|  |    370 | 
 | 
|  |    371 | fun the_preproc (Exec { preproc = Preproc x, ...}) = x;
 | 
|  |    372 | fun the_spec (Exec { spec = Spec x, ...}) = x;
 | 
|  |    373 | val the_funcs = #funcs o the_spec;
 | 
|  |    374 | val the_dcontrs = #dconstrs o the_spec;
 | 
|  |    375 | val the_dtyps = #dtyps o the_spec;
 | 
|  |    376 | val map_preproc = map_exec o apfst o map_preproc;
 | 
|  |    377 | val map_funcs = map_exec o apsnd o map_spec o apfst o apfst;
 | 
|  |    378 | val map_dconstrs = map_exec o apsnd o map_spec o apfst o apsnd;
 | 
|  |    379 | val map_dtyps = map_exec o apsnd o map_spec o apsnd;
 | 
|  |    380 | 
 | 
|  |    381 | 
 | 
|  |    382 | (** code data structures **)
 | 
|  |    383 | 
 | 
|  |    384 | (*private copy avoids potential conflict of table exceptions*)
 | 
|  |    385 | structure Datatab = TableFun(type key = int val ord = int_ord);
 | 
|  |    386 | 
 | 
|  |    387 | 
 | 
|  |    388 | (* data slots *)
 | 
|  |    389 | 
 | 
|  |    390 | local
 | 
|  |    391 | 
 | 
|  |    392 | type kind = {
 | 
|  |    393 |   name: string,
 | 
|  |    394 |   empty: Object.T,
 | 
|  |    395 |   merge: Pretty.pp -> Object.T * Object.T -> Object.T,
 | 
| 20937 |    396 |   purge: theory option -> CodegenConsts.const list option -> Object.T -> Object.T
 | 
| 20600 |    397 | };
 | 
|  |    398 | 
 | 
|  |    399 | val kinds = ref (Datatab.empty: kind Datatab.table);
 | 
|  |    400 | 
 | 
|  |    401 | fun invoke meth_name meth_fn k =
 | 
|  |    402 |   (case Datatab.lookup (! kinds) k of
 | 
|  |    403 |     SOME kind => meth_fn kind |> transform_failure (fn exn =>
 | 
|  |    404 |       EXCEPTION (exn, "Code data method " ^ #name kind ^ "." ^ meth_name ^ " failed"))
 | 
|  |    405 |   | NONE => sys_error ("Invalid code data identifier " ^ string_of_int k));
 | 
|  |    406 | 
 | 
|  |    407 | 
 | 
|  |    408 | in
 | 
|  |    409 | 
 | 
|  |    410 | fun invoke_name k   = invoke "name" (K o #name) k ();
 | 
|  |    411 | fun invoke_empty k  = invoke "empty" (K o #empty) k ();
 | 
|  |    412 | fun invoke_merge pp = invoke "merge" (fn kind => #merge kind pp);
 | 
| 20937 |    413 | fun invoke_purge thy_opt cs = invoke "purge" (fn kind => #purge kind thy_opt cs);
 | 
| 20600 |    414 | 
 | 
|  |    415 | fun declare name empty merge purge =
 | 
|  |    416 |   let
 | 
|  |    417 |     val k = serial ();
 | 
|  |    418 |     val kind = {name = name, empty = empty, merge = merge, purge = purge};
 | 
|  |    419 |     val _ = conditional (Datatab.exists (equal name o #name o #2) (! kinds)) (fn () =>
 | 
|  |    420 |       warning ("Duplicate declaration of code data " ^ quote name));
 | 
|  |    421 |     val _ = change kinds (Datatab.update (k, kind));
 | 
|  |    422 |   in k end;
 | 
|  |    423 | 
 | 
|  |    424 | end; (*local*)
 | 
|  |    425 | 
 | 
|  |    426 | 
 | 
|  |    427 | (* theory store *)
 | 
|  |    428 | 
 | 
|  |    429 | type data = Object.T Datatab.table;
 | 
|  |    430 | 
 | 
|  |    431 | structure CodeData = TheoryDataFun
 | 
|  |    432 | (struct
 | 
|  |    433 |   val name = "Pure/codegen_data";
 | 
|  |    434 |   type T = exec * data ref;
 | 
|  |    435 |   val empty = (empty_exec, ref Datatab.empty : data ref);
 | 
|  |    436 |   fun copy (exec, data) = (exec, ref (! data));
 | 
|  |    437 |   val extend = copy;
 | 
|  |    438 |   fun merge pp ((exec1, data1), (exec2, data2)) =
 | 
|  |    439 |     let
 | 
|  |    440 |       val (touched, exec) = merge_exec (exec1, exec2);
 | 
| 20937 |    441 |       val data1' = Datatab.map' (invoke_purge NONE touched) (! data1);
 | 
|  |    442 |       val data2' = Datatab.map' (invoke_purge NONE touched) (! data2);
 | 
| 20600 |    443 |       val data = Datatab.join (invoke_merge pp) (data1', data2');
 | 
|  |    444 |     in (exec, ref data) end;
 | 
|  |    445 |   fun print thy (exec, _) =
 | 
|  |    446 |     let
 | 
|  |    447 |       val ctxt = ProofContext.init thy;
 | 
|  |    448 |       fun pretty_func (s, lthms) =
 | 
|  |    449 |         (Pretty.block o Pretty.fbreaks) (
 | 
|  |    450 |           Pretty.str s :: pretty_sdthms ctxt lthms
 | 
|  |    451 |         );
 | 
|  |    452 |       fun pretty_dtyp (s, cos) =
 | 
|  |    453 |         (Pretty.block o Pretty.breaks) (
 | 
|  |    454 |           Pretty.str s
 | 
|  |    455 |           :: Pretty.str "="
 | 
|  |    456 |           :: Pretty.separate "|" (map (fn (c, []) => Pretty.str c
 | 
|  |    457 |                | (c, tys) =>
 | 
|  |    458 |                    Pretty.block
 | 
|  |    459 |                       (Pretty.str c :: Pretty.brk 1 :: Pretty.str "of" :: Pretty.brk 1
 | 
|  |    460 |                       :: Pretty.breaks (map (Pretty.quote o Sign.pretty_typ thy) tys))) cos)
 | 
|  |    461 |         )
 | 
|  |    462 |       val inlines = (#inlines o the_preproc) exec;
 | 
|  |    463 |       val funs = the_funcs exec
 | 
|  |    464 |         |> Consttab.dest
 | 
|  |    465 |         |> (map o apfst) (CodegenConsts.string_of_const thy)
 | 
|  |    466 |         |> sort (string_ord o pairself fst);
 | 
|  |    467 |       val dtyps = the_dtyps exec
 | 
|  |    468 |         |> Symtab.dest
 | 
|  |    469 |         |> map (fn (dtco, ((vs, cos), _)) => (Sign.string_of_typ thy (Type (dtco, map TFree vs)), cos))
 | 
|  |    470 |         |> sort (string_ord o pairself fst)
 | 
|  |    471 |     in
 | 
|  |    472 |       (Pretty.writeln o Pretty.block o Pretty.fbreaks) ([
 | 
|  |    473 |         Pretty.str "code theorems:",
 | 
|  |    474 |         Pretty.str "function theorems:" ] @
 | 
|  |    475 |           map pretty_func funs @ [
 | 
|  |    476 |         Pretty.block (
 | 
|  |    477 |           Pretty.str "inlined theorems:"
 | 
|  |    478 |           :: Pretty.fbrk
 | 
|  |    479 |           :: (Pretty.fbreaks o map (ProofContext.pretty_thm ctxt)) inlines
 | 
|  |    480 |         ),
 | 
|  |    481 |         Pretty.block (
 | 
|  |    482 |           Pretty.str "datatypes:"
 | 
|  |    483 |           :: Pretty.fbrk
 | 
|  |    484 |           :: (Pretty.fbreaks o map pretty_dtyp) dtyps
 | 
|  |    485 |         )]
 | 
|  |    486 |       )
 | 
|  |    487 |     end;
 | 
|  |    488 | end);
 | 
|  |    489 | 
 | 
|  |    490 | fun print_thms thy = CodeData.print thy;
 | 
|  |    491 | 
 | 
|  |    492 | fun init k = CodeData.map
 | 
|  |    493 |   (fn (exec, data) => (exec, ref (Datatab.update (k, invoke_empty k) (! data))));
 | 
|  |    494 | 
 | 
|  |    495 | fun get k dest data =
 | 
|  |    496 |   (case Datatab.lookup data k of
 | 
|  |    497 |     SOME x => (dest x handle Match =>
 | 
|  |    498 |       error ("Failed to access code data " ^ quote (invoke_name k)))
 | 
|  |    499 |   | NONE => error ("Uninitialized code data " ^ quote (invoke_name k)));
 | 
|  |    500 | 
 | 
|  |    501 | fun put k mk x = Datatab.update (k, mk x);
 | 
|  |    502 | 
 | 
| 20937 |    503 | fun map_exec_purge touched f thy =
 | 
| 20600 |    504 |   CodeData.map (fn (exec, data) => 
 | 
| 20937 |    505 |     (f exec, ref (Datatab.map' (invoke_purge (SOME thy) touched) (! data)))) thy;
 | 
| 20600 |    506 | 
 | 
|  |    507 | val get_exec = fst o CodeData.get;
 | 
|  |    508 | 
 | 
|  |    509 | val _ = Context.add_setup CodeData.init;
 | 
|  |    510 | 
 | 
|  |    511 | 
 | 
|  |    512 | 
 | 
|  |    513 | (** theorem transformation and certification **)
 | 
|  |    514 | 
 | 
|  |    515 | fun rewrite_func rewrites thm =
 | 
|  |    516 |   let
 | 
| 20844 |    517 |     val rewrite = Tactic.rewrite false rewrites;
 | 
|  |    518 |     val (ct_eq, [ct_lhs, ct_rhs]) = (Drule.strip_comb o Thm.cprop_of) thm;
 | 
|  |    519 |     val Const ("==", _) = Thm.term_of ct_eq;
 | 
| 20600 |    520 |     val (ct_f, ct_args) = Drule.strip_comb ct_lhs;
 | 
|  |    521 |     val rhs' = rewrite ct_rhs;
 | 
|  |    522 |     val args' = map rewrite ct_args;
 | 
|  |    523 |     val lhs' = Thm.symmetric (fold (fn th1 => fn th2 => Thm.combination th2 th1)
 | 
|  |    524 |       args' (Thm.reflexive ct_f));
 | 
|  |    525 |   in
 | 
|  |    526 |     Thm.transitive (Thm.transitive lhs' thm) rhs'
 | 
|  |    527 |   end handle Bind => raise ERROR "rewrite_func"
 | 
|  |    528 | 
 | 
|  |    529 | fun common_typ_funcs thy [] = []
 | 
|  |    530 |   | common_typ_funcs thy [thm] = [thm]
 | 
|  |    531 |   | common_typ_funcs thy thms =
 | 
|  |    532 |       let
 | 
|  |    533 |         fun incr_thm thm max =
 | 
|  |    534 |           let
 | 
|  |    535 |             val thm' = incr_indexes max thm;
 | 
|  |    536 |             val max' = (maxidx_of_typ o fastype_of o Drule.plain_prop_of) thm' + 1;
 | 
|  |    537 |           in (thm', max') end;
 | 
|  |    538 |         val (thms', maxidx) = fold_map incr_thm thms 0;
 | 
|  |    539 |         val (ty1::tys) = map (typ_func thy) thms;
 | 
|  |    540 |         fun unify ty env = Sign.typ_unify thy (ty1, ty) env
 | 
|  |    541 |           handle Type.TUNIFY =>
 | 
|  |    542 |             error ("Type unificaton failed, while unifying function equations\n"
 | 
|  |    543 |             ^ (cat_lines o map Display.string_of_thm) thms
 | 
|  |    544 |             ^ "\nwith types\n"
 | 
|  |    545 |             ^ (cat_lines o map (Sign.string_of_typ thy)) (ty1 :: tys));
 | 
|  |    546 |         val (env, _) = fold unify tys (Vartab.empty, maxidx)
 | 
|  |    547 |         val instT = Vartab.fold (fn (x_i, (sort, ty)) =>
 | 
|  |    548 |           cons (Thm.ctyp_of thy (TVar (x_i, sort)), Thm.ctyp_of thy ty)) env [];
 | 
|  |    549 |       in map (Thm.instantiate (instT, [])) thms end;
 | 
|  |    550 | 
 | 
| 20844 |    551 | fun certify_const thy c c_thms =
 | 
| 20600 |    552 |   let
 | 
|  |    553 |     fun cert (c', thm) = if CodegenConsts.eq_const (c, c')
 | 
|  |    554 |       then thm else bad_thm ("Wrong head of function equation,\nexpected constant "
 | 
|  |    555 |         ^ CodegenConsts.string_of_const thy c) thm
 | 
| 20844 |    556 |   in map cert c_thms end;
 | 
| 20600 |    557 | 
 | 
|  |    558 | fun mk_cos tyco vs cos =
 | 
|  |    559 |   let
 | 
|  |    560 |     val dty = Type (tyco, map TFree vs);
 | 
|  |    561 |     fun mk_co (co, tys) = (Const (co, (tys ---> dty)), map I tys);
 | 
|  |    562 |   in map mk_co cos end;
 | 
|  |    563 | 
 | 
|  |    564 | fun mk_co_args (co, tys) ctxt =
 | 
|  |    565 |   let
 | 
|  |    566 |     val names = Name.invents ctxt "a" (length tys);
 | 
|  |    567 |     val ctxt' = fold Name.declare names ctxt;
 | 
|  |    568 |     val vs = map2 (fn v => fn ty => Free (fst (v, 0), I ty)) names tys;
 | 
|  |    569 |   in (vs, ctxt') end;
 | 
|  |    570 | 
 | 
|  |    571 | fun check_freeness thy cos thms =
 | 
|  |    572 |   let
 | 
|  |    573 |     val props = AList.make Drule.plain_prop_of thms;
 | 
|  |    574 |     fun sym_product [] = []
 | 
|  |    575 |       | sym_product (x::xs) = map (pair x) xs @ sym_product xs;
 | 
|  |    576 |     val quodlibet =
 | 
|  |    577 |       let
 | 
|  |    578 |         val judg = ObjectLogic.fixed_judgment (the_context ()) "x";
 | 
|  |    579 |         val [free] = fold_aterms (fn v as Free _ => cons v | _ => I) judg [];
 | 
|  |    580 |         val judg' = Term.subst_free [(free, Bound 0)] judg;
 | 
|  |    581 |         val prop = Type ("prop", []);
 | 
|  |    582 |         val prop' = fastype_of judg';
 | 
|  |    583 |       in
 | 
|  |    584 |         Const ("all", (prop' --> prop) --> prop) $ Abs ("P", prop', judg')
 | 
|  |    585 |       end;
 | 
|  |    586 |     fun check_inj (co, []) =
 | 
|  |    587 |           NONE
 | 
|  |    588 |       | check_inj (co, tys) =
 | 
|  |    589 |           let
 | 
|  |    590 |             val ((xs, ys), _) = Name.context
 | 
|  |    591 |               |> mk_co_args (co, tys)
 | 
|  |    592 |               ||>> mk_co_args (co, tys);
 | 
|  |    593 |             val prem = Logic.mk_equals
 | 
|  |    594 |               (list_comb (co, xs), list_comb (co, ys));
 | 
|  |    595 |             val concl = Logic.mk_conjunction_list
 | 
|  |    596 |               (map2 (curry Logic.mk_equals) xs ys);
 | 
|  |    597 |             val t = Logic.mk_implies (prem, concl);
 | 
|  |    598 |           in case find_first (curry Term.could_unify t o snd) props
 | 
|  |    599 |            of SOME (thm, _) => SOME thm
 | 
|  |    600 |             | NONE => error ("Could not prove injectiveness statement\n"
 | 
|  |    601 |                ^ Sign.string_of_term thy t
 | 
|  |    602 |                ^ "\nfor constructor "
 | 
|  |    603 |                ^ CodegenConsts.string_of_const_typ thy (dest_Const co)
 | 
|  |    604 |                ^ "\nwith theorems\n" ^ cat_lines (map string_of_thm thms))
 | 
|  |    605 |           end;
 | 
|  |    606 |     fun check_dist ((co1, tys1), (co2, tys2)) =
 | 
|  |    607 |           let
 | 
|  |    608 |             val ((xs1, xs2), _) = Name.context
 | 
|  |    609 |               |> mk_co_args (co1, tys1)
 | 
|  |    610 |               ||>> mk_co_args (co2, tys2);
 | 
|  |    611 |             val prem = Logic.mk_equals
 | 
|  |    612 |               (list_comb (co1, xs1), list_comb (co2, xs2));
 | 
|  |    613 |             val t = Logic.mk_implies (prem, quodlibet);
 | 
|  |    614 |           in case find_first (curry Term.could_unify t o snd) props
 | 
|  |    615 |            of SOME (thm, _) => thm
 | 
|  |    616 |             | NONE => error ("Could not prove distinctness statement\n"
 | 
|  |    617 |                ^ Sign.string_of_term thy t
 | 
|  |    618 |                ^ "\nfor constructors "
 | 
|  |    619 |                ^ CodegenConsts.string_of_const_typ thy (dest_Const co1)
 | 
|  |    620 |                ^ " and "
 | 
|  |    621 |                ^ CodegenConsts.string_of_const_typ thy (dest_Const co2)
 | 
|  |    622 |                ^ "\nwith theorems\n" ^ cat_lines (map string_of_thm thms))
 | 
|  |    623 |           end;
 | 
|  |    624 |   in (map_filter check_inj cos, map check_dist (sym_product cos)) end;
 | 
|  |    625 | 
 | 
|  |    626 | fun certify_datatype thy dtco cs thms =
 | 
|  |    627 |   (op @) (check_freeness thy cs thms);
 | 
|  |    628 | 
 | 
|  |    629 | 
 | 
|  |    630 | 
 | 
|  |    631 | (** interfaces **)
 | 
|  |    632 | 
 | 
|  |    633 | fun add_func thm thy =
 | 
|  |    634 |   let
 | 
|  |    635 |     val thms = mk_func thy thm;
 | 
|  |    636 |     val cs = map fst thms;
 | 
|  |    637 |   in
 | 
|  |    638 |     map_exec_purge (SOME cs) (map_funcs 
 | 
|  |    639 |      (fold (fn (c, thm) => Consttab.map_default
 | 
|  |    640 |        (c, (Susp.value [], [])) (add_thm thm)) thms)) thy
 | 
|  |    641 |   end;
 | 
|  |    642 | 
 | 
|  |    643 | fun del_func thm thy =
 | 
|  |    644 |   let
 | 
|  |    645 |     val thms = mk_func thy thm;
 | 
|  |    646 |     val cs = map fst thms;
 | 
|  |    647 |   in
 | 
|  |    648 |     map_exec_purge (SOME cs) (map_funcs
 | 
|  |    649 |      (fold (fn (c, thm) => Consttab.map_entry c
 | 
|  |    650 |        (del_thm thm)) thms)) thy
 | 
|  |    651 |   end;
 | 
|  |    652 | 
 | 
|  |    653 | fun add_funcl (c, lthms) thy =
 | 
|  |    654 |   let
 | 
|  |    655 |     val c' = CodegenConsts.norm thy c;
 | 
| 20844 |    656 |     val lthms' = certificate thy (fn thy => certify_const thy c' o maps (mk_func thy)) lthms;
 | 
| 20600 |    657 |   in
 | 
|  |    658 |     map_exec_purge (SOME [c]) (map_funcs (Consttab.map_default (c', (Susp.value [], []))
 | 
|  |    659 |       (add_lthms lthms'))) thy
 | 
|  |    660 |   end;
 | 
|  |    661 | 
 | 
|  |    662 | fun add_datatype (tyco, (vs_cos as (vs, cos), lthms)) thy =
 | 
|  |    663 |   let
 | 
|  |    664 |     val cs = mk_cos tyco vs cos;
 | 
|  |    665 |     val consts = map (CodegenConsts.norm_of_typ thy o dest_Const o fst) cs;
 | 
|  |    666 |     val add =
 | 
|  |    667 |       map_dtyps (Symtab.update_new (tyco,
 | 
| 20844 |    668 |         (vs_cos, certificate thy (fn thy => certify_datatype thy tyco cs) lthms)))
 | 
| 20600 |    669 |       #> map_dconstrs (fold (fn c => Consttab.update (c, tyco)) consts)
 | 
|  |    670 |   in map_exec_purge (SOME consts) add thy end;
 | 
|  |    671 | 
 | 
|  |    672 | fun del_datatype tyco thy =
 | 
|  |    673 |   let
 | 
| 20639 |    674 |     val SOME ((vs, cos), _) = Symtab.lookup ((the_dtyps o get_exec) thy) tyco;
 | 
|  |    675 |     val cs = mk_cos tyco vs cos;
 | 
|  |    676 |     val consts = map (CodegenConsts.norm_of_typ thy o dest_Const o fst) cs;
 | 
| 20600 |    677 |     val del =
 | 
|  |    678 |       map_dtyps (Symtab.delete tyco)
 | 
| 20639 |    679 |       #> map_dconstrs (fold Consttab.delete consts)
 | 
|  |    680 |   in map_exec_purge (SOME consts) del thy end;
 | 
| 20600 |    681 | 
 | 
|  |    682 | fun add_inline thm thy =
 | 
| 20844 |    683 |   (map_exec_purge NONE o map_preproc o apfst o apfst) (fold (insert eq_thm) (mk_rew thy thm)) thy;
 | 
| 20600 |    684 | 
 | 
|  |    685 | fun del_inline thm thy =
 | 
| 20844 |    686 |   (map_exec_purge NONE o map_preproc o apfst o apfst) (fold (remove eq_thm) (mk_rew thy thm)) thy ;
 | 
|  |    687 | 
 | 
|  |    688 | fun add_inline_proc f =
 | 
|  |    689 |   (map_exec_purge NONE o map_preproc o apfst o apsnd) (cons (serial (), f));
 | 
|  |    690 | 
 | 
|  |    691 | fun add_constrains f =
 | 
|  |    692 |   (map_exec_purge NONE o map_preproc o apsnd o apfst) (cons (serial (), f));
 | 
| 20600 |    693 | 
 | 
|  |    694 | fun add_preproc f =
 | 
| 20844 |    695 |   (map_exec_purge NONE o map_preproc o apsnd o apsnd) (cons (serial (), f));
 | 
|  |    696 | 
 | 
|  |    697 | local
 | 
|  |    698 | 
 | 
|  |    699 | fun gen_apply_constrain prep post const_typ thy fs x =
 | 
|  |    700 |   let
 | 
|  |    701 |     val ts = prep x;
 | 
|  |    702 |     val tvars = (fold o fold_aterms) Term.add_tvars ts [];
 | 
|  |    703 |     val consts = (fold o fold_aterms) (fn Const c => cons c | _ => I) ts [];
 | 
|  |    704 |     fun insts_of const_typ (c, ty) =
 | 
|  |    705 |       let
 | 
|  |    706 |         val ty_decl = const_typ (c, ty);
 | 
|  |    707 |         val env = Vartab.dest (Type.raw_match (ty_decl, ty) Vartab.empty);
 | 
|  |    708 |         val insts = map_filter
 | 
|  |    709 |          (fn (v, (sort, TVar (_, sort'))) =>
 | 
|  |    710 |                 if Sorts.sort_le (Sign.classes_of thy) (sort, sort')
 | 
|  |    711 |                 then NONE else SOME (v, sort)
 | 
|  |    712 |            | _ => NONE) env
 | 
|  |    713 |       in 
 | 
|  |    714 |         insts
 | 
|  |    715 |       end
 | 
|  |    716 |     val const_insts = case const_typ
 | 
|  |    717 |      of NONE => []
 | 
|  |    718 |       | SOME const_typ => maps (insts_of const_typ) consts;
 | 
|  |    719 |     fun add_inst (v, sort') =
 | 
|  |    720 |       let
 | 
|  |    721 |         val sort = (the o AList.lookup (op =) tvars) v
 | 
|  |    722 |       in
 | 
|  |    723 |         AList.map_default (op =) (v, (sort, sort))
 | 
|  |    724 |           (apsnd (fn sort => Sorts.inter_sort (Sign.classes_of thy) (sort, sort')))
 | 
|  |    725 |       end;
 | 
|  |    726 |     val inst =
 | 
|  |    727 |       []
 | 
|  |    728 |       |> fold (fn f => fold add_inst (f thy ts)) fs
 | 
|  |    729 |       |> fold add_inst const_insts;
 | 
|  |    730 |   in
 | 
|  |    731 |     post thy inst x
 | 
|  |    732 |   end;
 | 
| 20600 |    733 | 
 | 
| 20844 |    734 | val apply_constrain = gen_apply_constrain (maps
 | 
|  |    735 |   ((fn (args, rhs) => rhs :: (snd o strip_comb) args) o Logic.dest_equals o Thm.prop_of))
 | 
|  |    736 |   (fn thy => fn inst => map (check_typ_classop thy o Thm.instantiate (map (fn (v, (sort, sort')) =>
 | 
|  |    737 |     (Thm.ctyp_of thy (TVar (v, sort)), Thm.ctyp_of thy (TVar (v, sort')))
 | 
|  |    738 |   ) inst, []))) NONE;
 | 
|  |    739 | fun apply_constrain_cterm thy const_typ = gen_apply_constrain (single o Thm.term_of)
 | 
|  |    740 |   (fn thy => fn inst => pair inst o Thm.cterm_of thy o map_types
 | 
|  |    741 |     (TermSubst.instantiateT (map (fn (v, (sort, sort')) => ((v, sort), TVar (v, sort'))) inst)) o Thm.term_of) (SOME const_typ) thy;
 | 
|  |    742 | 
 | 
|  |    743 | fun gen_apply_inline_proc prep post thy f x =
 | 
|  |    744 |   let
 | 
|  |    745 |     val cts = prep x;
 | 
|  |    746 |     val rews = map (check_rew thy) (f thy cts);
 | 
|  |    747 |   in post rews x end;
 | 
|  |    748 | 
 | 
|  |    749 | val apply_inline_proc = gen_apply_inline_proc (maps
 | 
|  |    750 |   ((fn [args, rhs] => rhs :: (snd o Drule.strip_comb) args) o snd o Drule.strip_comb o Thm.cprop_of))
 | 
|  |    751 |   (fn rews => map (rewrite_func rews));
 | 
|  |    752 | val apply_inline_proc_cterm = gen_apply_inline_proc single
 | 
|  |    753 |   (Tactic.rewrite false);
 | 
| 20600 |    754 | 
 | 
| 20844 |    755 | fun apply_preproc thy f [] = []
 | 
|  |    756 |   | apply_preproc thy f (thms as (thm :: _)) =
 | 
|  |    757 |       let
 | 
|  |    758 |         val thms' = f thy thms;
 | 
|  |    759 |         val c = (CodegenConsts.norm_of_typ thy o fst o dest_func thy) thm;
 | 
|  |    760 |       in (certify_const thy c o map (mk_head thy)) thms' end;
 | 
|  |    761 | 
 | 
|  |    762 | fun cmp_thms thy =
 | 
|  |    763 |   make_ord (fn (thm1, thm2) => not (Sign.typ_instance thy (typ_func thy thm1, typ_func thy thm2)));
 | 
|  |    764 | 
 | 
|  |    765 | fun rhs_conv conv thm =
 | 
|  |    766 |   let
 | 
|  |    767 |     val thm' = (conv o snd o Drule.dest_equals o Thm.cprop_of) thm;
 | 
|  |    768 |   in Thm.transitive thm thm' end
 | 
|  |    769 | 
 | 
|  |    770 | fun drop_classes thy inst thm =
 | 
|  |    771 |   let
 | 
|  |    772 |     val unconstr = map (fn (v, (_, sort')) =>
 | 
|  |    773 |       (Thm.ctyp_of thy o TVar) (v, sort')) inst;
 | 
|  |    774 |     val instmap = map (fn (v, (sort, _)) =>
 | 
|  |    775 |       pairself (Thm.ctyp_of thy o TVar) ((v, []), (v, sort))) inst;
 | 
|  |    776 |   in
 | 
|  |    777 |     thm
 | 
|  |    778 |     |> fold Thm.unconstrainT unconstr
 | 
|  |    779 |     |> Thm.instantiate (instmap, [])
 | 
|  |    780 |     |> Tactic.rule_by_tactic ((REPEAT o CHANGED o ALLGOALS o Tactic.resolve_tac) (AxClass.class_intros thy))
 | 
|  |    781 |   end;
 | 
|  |    782 | 
 | 
|  |    783 | in
 | 
| 20600 |    784 | 
 | 
|  |    785 | fun preprocess thy thms =
 | 
| 20844 |    786 |   thms
 | 
|  |    787 |   |> fold (fn (_, f) => apply_preproc thy f) ((#preprocs o the_preproc o get_exec) thy)
 | 
|  |    788 |   |> map (rewrite_func ((#inlines o the_preproc o get_exec) thy))
 | 
|  |    789 |   |> apply_constrain thy ((map snd o #constrains o the_preproc o get_exec) thy)
 | 
|  |    790 |   |> map (rewrite_func ((#inlines o the_preproc o get_exec) thy))
 | 
|  |    791 |   |> fold (fn (_, f) => apply_inline_proc thy f) ((#inline_procs o the_preproc o get_exec) thy)
 | 
|  |    792 |   |> map (snd o check_func false thy)
 | 
|  |    793 |   |> sort (cmp_thms thy)
 | 
|  |    794 |   |> common_typ_funcs thy;
 | 
| 20600 |    795 | 
 | 
| 20844 |    796 | fun preprocess_cterm thy const_typ ct =
 | 
|  |    797 |   ct
 | 
|  |    798 |   |> apply_constrain_cterm thy const_typ ((map snd o #constrains o the_preproc o get_exec) thy)
 | 
|  |    799 |   |-> (fn inst =>
 | 
|  |    800 |      Thm.reflexive
 | 
|  |    801 |   #> fold (rhs_conv o Tactic.rewrite false o single) ((#inlines o the_preproc o get_exec) thy)
 | 
|  |    802 |   #> fold (fn (_, f) => rhs_conv (apply_inline_proc_cterm thy f)) ((#inline_procs o the_preproc o get_exec) thy)
 | 
|  |    803 |   #> (fn thm => (drop_classes thy inst thm, ((fn xs => nth xs 1) o snd o Drule.strip_comb o Thm.cprop_of) thm))
 | 
|  |    804 |   );
 | 
|  |    805 | 
 | 
|  |    806 | end; (*local*)
 | 
| 20600 |    807 | 
 | 
|  |    808 | fun these_funcs thy c =
 | 
|  |    809 |   let
 | 
| 20844 |    810 |     val funcs_1 =
 | 
| 20600 |    811 |       Consttab.lookup ((the_funcs o get_exec) thy) c
 | 
|  |    812 |       |> Option.map (Susp.force o fst)
 | 
|  |    813 |       |> these
 | 
|  |    814 |       |> map (Thm.transfer thy);
 | 
| 20844 |    815 |     val funcs_2 = case funcs_1
 | 
|  |    816 |      of [] => get_prim_def_funcs thy c
 | 
|  |    817 |       | xs => xs;
 | 
| 20600 |    818 |     fun drop_refl thy = filter_out (is_equal o Term.fast_term_ord o Logic.dest_equals
 | 
|  |    819 |       o ObjectLogic.drop_judgment thy o Drule.plain_prop_of);
 | 
|  |    820 |   in
 | 
| 20844 |    821 |     funcs_2
 | 
| 20600 |    822 |     |> preprocess thy
 | 
|  |    823 |     |> drop_refl thy
 | 
|  |    824 |   end;
 | 
|  |    825 | 
 | 
|  |    826 | fun get_datatype thy tyco =
 | 
|  |    827 |   Symtab.lookup ((the_dtyps o get_exec) thy) tyco
 | 
|  |    828 |   |> Option.map (fn (spec, thms) => (Susp.force thms; spec));
 | 
|  |    829 | 
 | 
|  |    830 | fun get_datatype_of_constr thy c =
 | 
|  |    831 |   Consttab.lookup ((the_dcontrs o get_exec) thy) c
 | 
|  |    832 |   |> (Option.map o tap) (fn dtco => get_datatype thy dtco);
 | 
|  |    833 | 
 | 
| 20937 |    834 | fun typ_funcs thy (c as (name, _), []) = (case AxClass.class_of_param thy name
 | 
|  |    835 |      of SOME class => CodegenConsts.disc_typ_of_classop thy c
 | 
|  |    836 |       | NONE => (case Option.map (Susp.force o fst) (Consttab.lookup ((the_funcs o get_exec) thy) c)
 | 
|  |    837 |          of SOME [eq] => typ_func thy eq
 | 
|  |    838 |           | _ => Sign.the_const_type thy name))
 | 
|  |    839 |   | typ_funcs thy (_, eq :: _) = typ_func thy eq;
 | 
| 20600 |    840 | 
 | 
|  |    841 | 
 | 
|  |    842 | (** code attributes **)
 | 
|  |    843 | 
 | 
|  |    844 | local
 | 
|  |    845 |   fun add_simple_attribute (name, f) =
 | 
|  |    846 |     (Codegen.add_attribute name o (Scan.succeed o Thm.declaration_attribute))
 | 
| 20897 |    847 |       (fn th => Context.mapping (f th) I);
 | 
| 20600 |    848 | in
 | 
|  |    849 |   val _ = map (Context.add_setup o add_simple_attribute) [
 | 
|  |    850 |     ("func", add_func),
 | 
|  |    851 |     ("nofunc", del_func),
 | 
|  |    852 |     ("unfold", (fn thm => Codegen.add_unfold thm #> add_inline thm)),
 | 
|  |    853 |     ("inline", add_inline),
 | 
|  |    854 |     ("noinline", del_inline)
 | 
|  |    855 |   ]
 | 
|  |    856 | end; (*local*)
 | 
|  |    857 | 
 | 
|  |    858 | end; (*struct*)
 | 
|  |    859 | 
 | 
|  |    860 | 
 | 
|  |    861 | 
 | 
|  |    862 | (** type-safe interfaces for data depedent on executable content **)
 | 
|  |    863 | 
 | 
|  |    864 | signature CODE_DATA_ARGS =
 | 
|  |    865 | sig
 | 
|  |    866 |   val name: string
 | 
|  |    867 |   type T
 | 
|  |    868 |   val empty: T
 | 
|  |    869 |   val merge: Pretty.pp -> T * T -> T
 | 
| 20937 |    870 |   val purge: theory option -> CodegenConsts.const list option -> T -> T
 | 
| 20600 |    871 | end;
 | 
|  |    872 | 
 | 
|  |    873 | signature CODE_DATA =
 | 
|  |    874 | sig
 | 
|  |    875 |   type T
 | 
|  |    876 |   val init: theory -> theory
 | 
|  |    877 |   val get: theory -> T
 | 
|  |    878 |   val change: theory -> (T -> T) -> T
 | 
|  |    879 |   val change_yield: theory -> (T -> 'a * T) -> 'a * T
 | 
|  |    880 | end;
 | 
|  |    881 | 
 | 
|  |    882 | functor CodeDataFun(Data: CODE_DATA_ARGS): CODE_DATA =
 | 
|  |    883 | struct
 | 
|  |    884 | 
 | 
|  |    885 | type T = Data.T;
 | 
|  |    886 | exception Data of T;
 | 
|  |    887 | fun dest (Data x) = x
 | 
|  |    888 | 
 | 
|  |    889 | val kind = CodegenData.declare Data.name (Data Data.empty)
 | 
|  |    890 |   (fn pp => fn (Data x1, Data x2) => Data (Data.merge pp (x1, x2)))
 | 
| 20937 |    891 |   (fn thy_opt => fn cs => fn Data x => Data (Data.purge thy_opt cs x));
 | 
| 20600 |    892 | 
 | 
|  |    893 | val init = CodegenData.init kind;
 | 
|  |    894 | fun get thy = CodegenData.get kind dest ((! o snd o CodegenData.CodeData.get) thy);
 | 
|  |    895 | fun change thy f =
 | 
|  |    896 |   let
 | 
|  |    897 |     val data_ref = (snd o CodegenData.CodeData.get) thy;
 | 
|  |    898 |     val x = (f o CodegenData.get kind dest o !) data_ref;
 | 
|  |    899 |     val data = CodegenData.put kind Data x (! data_ref);
 | 
|  |    900 |   in (data_ref := data; x) end;
 | 
|  |    901 | fun change_yield thy f =
 | 
|  |    902 |   let
 | 
|  |    903 |     val data_ref = (snd o CodegenData.CodeData.get) thy;
 | 
|  |    904 |     val (y, x) = (f o CodegenData.get kind dest o !) data_ref;
 | 
|  |    905 |     val data = CodegenData.put kind Data x (! data_ref);
 | 
|  |    906 |   in (data_ref := data; (y, x)) end;
 | 
|  |    907 | 
 | 
|  |    908 | end;
 | 
|  |    909 | 
 | 
|  |    910 | structure CodegenData : CODEGEN_DATA =
 | 
|  |    911 | struct
 | 
|  |    912 | 
 | 
|  |    913 | open CodegenData;
 | 
|  |    914 | 
 | 
|  |    915 | end;
 |