| author | wenzelm | 
| Sun, 15 Nov 2009 15:14:02 +0100 | |
| changeset 33696 | 2c7c79ca6c23 | 
| parent 33531 | 2c0a4adbcf13 | 
| child 33708 | b45d3b8cc74e | 
| permissions | -rw-r--r-- | 
| 24219 | 1 | (* Title: Pure/Isar/code.ML | 
| 2 | Author: Florian Haftmann, TU Muenchen | |
| 3 | ||
| 31962 | 4 | Abstract executable code of theory. Management of data dependent on | 
| 5 | executable code. Cache assumes non-concurrent processing of a single theory. | |
| 24219 | 6 | *) | 
| 7 | ||
| 8 | signature CODE = | |
| 9 | sig | |
| 31957 | 10 | (*constants*) | 
| 11 | val check_const: theory -> term -> string | |
| 12 | val read_bare_const: theory -> string -> string * typ | |
| 13 | val read_const: theory -> string -> string | |
| 31962 | 14 | val string_of_const: theory -> string -> string | 
| 15 | val args_number: theory -> string -> int | |
| 31957 | 16 | |
| 31156 | 17 | (*constructor sets*) | 
| 18 | val constrset_of_consts: theory -> (string * typ) list | |
| 19 | -> string * ((string * sort) list * (string * typ list) list) | |
| 20 | ||
| 21 | (*code equations*) | |
| 31962 | 22 | val mk_eqn: theory -> thm * bool -> thm * bool | 
| 23 | val mk_eqn_warning: theory -> thm -> (thm * bool) option | |
| 24 | val mk_eqn_liberal: theory -> thm -> (thm * bool) option | |
| 31156 | 25 | val assert_eqn: theory -> thm * bool -> thm * bool | 
| 26 | val assert_eqns_const: theory -> string | |
| 27 | -> (thm * bool) list -> (thm * bool) list | |
| 31957 | 28 | val const_typ_eqn: theory -> thm -> string * typ | 
| 31156 | 29 | val typscheme_eqn: theory -> thm -> (string * sort) list * typ | 
| 32873 | 30 | val typscheme_eqns: theory -> string -> thm list -> (string * sort) list * typ | 
| 32929 | 31 | val standard_typscheme: theory -> thm list -> thm list | 
| 31156 | 32 | |
| 31962 | 33 | (*executable code*) | 
| 31156 | 34 | val add_datatype: (string * typ) list -> theory -> theory | 
| 35 | val add_datatype_cmd: string list -> theory -> theory | |
| 36 | val type_interpretation: | |
| 37 | (string * ((string * sort) list * (string * typ list) list) | |
| 38 | -> theory -> theory) -> theory -> theory | |
| 28368 | 39 | val add_eqn: thm -> theory -> theory | 
| 31088 | 40 | val add_nbe_eqn: thm -> theory -> theory | 
| 28368 | 41 | val add_default_eqn: thm -> theory -> theory | 
| 28703 | 42 | val add_default_eqn_attribute: attribute | 
| 43 | val add_default_eqn_attrib: Attrib.src | |
| 28368 | 44 | val del_eqn: thm -> theory -> theory | 
| 45 | val del_eqns: string -> theory -> theory | |
| 24844 
98c006a30218
certificates for code generator case expressions
 haftmann parents: 
24837diff
changeset | 46 | val add_case: thm -> theory -> theory | 
| 
98c006a30218
certificates for code generator case expressions
 haftmann parents: 
24837diff
changeset | 47 | val add_undefined: string -> theory -> theory | 
| 24219 | 48 | val get_datatype: theory -> string -> ((string * sort) list * (string * typ list) list) | 
| 24423 
ae9cd0e92423
overloaded definitions accompanied by explicit constants
 haftmann parents: 
24283diff
changeset | 49 | val get_datatype_of_constr: theory -> string -> string option | 
| 31156 | 50 | val these_eqns: theory -> string -> (thm * bool) list | 
| 31642 | 51 | val all_eqns: theory -> (thm * bool) list | 
| 30023 
55954f726043
permissive check for pattern discipline in case schemes
 haftmann parents: 
29970diff
changeset | 52 | val get_case_scheme: theory -> string -> (int * (int * string list)) option | 
| 31890 
e943b039f0ac
an intermediate step towards a refined translation of cases
 haftmann parents: 
31642diff
changeset | 53 | val undefineds: theory -> string list | 
| 24219 | 54 | val print_codesetup: theory -> unit | 
| 31957 | 55 | |
| 56 | (*infrastructure*) | |
| 31998 
2c7a24f74db9
code attributes use common underscore convention
 haftmann parents: 
31971diff
changeset | 57 | val set_code_target_attr: (string -> thm -> theory -> theory) -> theory -> theory | 
| 31957 | 58 | val purge_data: theory -> theory | 
| 24219 | 59 | end; | 
| 60 | ||
| 61 | signature CODE_DATA_ARGS = | |
| 62 | sig | |
| 63 | type T | |
| 64 | val empty: T | |
| 27609 | 65 | val purge: theory -> string list -> T -> T | 
| 24219 | 66 | end; | 
| 67 | ||
| 68 | signature CODE_DATA = | |
| 69 | sig | |
| 70 | type T | |
| 71 | val get: theory -> T | |
| 72 | val change: theory -> (T -> T) -> T | |
| 73 | val change_yield: theory -> (T -> 'a * T) -> 'a * T | |
| 74 | end; | |
| 75 | ||
| 76 | signature PRIVATE_CODE = | |
| 77 | sig | |
| 78 | include CODE | |
| 27609 | 79 | val declare_data: Object.T -> (theory -> string list -> Object.T -> Object.T) | 
| 80 | -> serial | |
| 24219 | 81 |   val get_data: serial * ('a -> Object.T) * (Object.T -> 'a)
 | 
| 82 | -> theory -> 'a | |
| 83 |   val change_data: serial * ('a -> Object.T) * (Object.T -> 'a)
 | |
| 84 |     -> theory -> ('a -> 'a) -> 'a
 | |
| 85 |   val change_yield_data: serial * ('a -> Object.T) * (Object.T -> 'a)
 | |
| 86 |     -> theory -> ('a -> 'b * 'a) -> 'b * 'a
 | |
| 87 | end; | |
| 88 | ||
| 89 | structure Code : PRIVATE_CODE = | |
| 90 | struct | |
| 91 | ||
| 31962 | 92 | (** auxiliary **) | 
| 93 | ||
| 94 | (* printing *) | |
| 31156 | 95 | |
| 32966 
5b21661fe618
indicate CRITICAL nature of various setmp combinators;
 wenzelm parents: 
32929diff
changeset | 96 | fun string_of_typ thy = setmp_CRITICAL show_sorts true (Syntax.string_of_typ_global thy); | 
| 31962 | 97 | |
| 31156 | 98 | fun string_of_const thy c = case AxClass.inst_of_param thy c | 
| 99 | of SOME (c, tyco) => Sign.extern_const thy c ^ " " ^ enclose "[" "]" (Sign.extern_type thy tyco) | |
| 100 | | NONE => Sign.extern_const thy c; | |
| 101 | ||
| 31962 | 102 | |
| 103 | (* constants *) | |
| 31156 | 104 | |
| 31962 | 105 | fun check_bare_const thy t = case try dest_Const t | 
| 106 | of SOME c_ty => c_ty | |
| 107 |   | NONE => error ("Not a constant: " ^ Syntax.string_of_term_global thy t);
 | |
| 31156 | 108 | |
| 31962 | 109 | fun check_const thy = AxClass.unoverload_const thy o check_bare_const thy; | 
| 110 | ||
| 111 | fun read_bare_const thy = check_bare_const thy o Syntax.read_term_global thy; | |
| 112 | ||
| 113 | fun read_const thy = AxClass.unoverload_const thy o read_bare_const thy; | |
| 31156 | 114 | |
| 32872 
019201eb7e07
variables in type schemes must be renamed simultaneously with variables in equations
 haftmann parents: 
32738diff
changeset | 115 | |
| 31962 | 116 | |
| 117 | (** data store **) | |
| 118 | ||
| 119 | (* code equations *) | |
| 120 | ||
| 33006 | 121 | type eqns = bool * (thm * bool) list; | 
| 122 | (*default flag, theorems with proper flag *) | |
| 31962 | 123 | |
| 124 | fun add_drop_redundant thy (thm, proper) thms = | |
| 125 | let | |
| 126 | val args_of = snd o strip_comb o map_types Type.strip_sorts | |
| 127 | o fst o Logic.dest_equals o Thm.plain_prop_of; | |
| 128 | val args = args_of thm; | |
| 129 | val incr_idx = Logic.incr_indexes ([], Thm.maxidx_of thm + 1); | |
| 130 | fun matches_args args' = length args <= length args' andalso | |
| 131 | Pattern.matchess thy (args, (map incr_idx o curry Library.take (length args)) args'); | |
| 132 | fun drop (thm', proper') = if (proper orelse not proper') | |
| 133 | andalso matches_args (args_of thm') then | |
| 32091 
30e2ffbba718
proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
 wenzelm parents: 
32070diff
changeset | 134 |         (warning ("Code generator: dropping redundant code equation\n" ^
 | 
| 
30e2ffbba718
proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
 wenzelm parents: 
32070diff
changeset | 135 | Display.string_of_thm_global thy thm'); true) | 
| 31962 | 136 | else false; | 
| 137 | in (thm, proper) :: filter_out drop thms end; | |
| 138 | ||
| 33006 | 139 | fun add_thm thy _ thm (false, thms) = (false, add_drop_redundant thy thm thms) | 
| 140 | | add_thm thy true thm (true, thms) = (true, thms @ [thm]) | |
| 141 | | add_thm thy false thm (true, thms) = (false, [thm]); | |
| 31962 | 142 | |
| 33006 | 143 | fun del_thm thm = apsnd (remove (eq_fst Thm.eq_thm_prop) (thm, true)); | 
| 31962 | 144 | |
| 145 | ||
| 146 | (* executable code data *) | |
| 147 | ||
| 148 | datatype spec = Spec of {
 | |
| 149 | history_concluded: bool, | |
| 150 | eqns: ((bool * eqns) * (serial * eqns) list) Symtab.table | |
| 151 | (*with explicit history*), | |
| 152 | dtyps: ((serial * ((string * sort) list * (string * typ list) list)) list) Symtab.table | |
| 153 | (*with explicit history*), | |
| 154 | cases: (int * (int * string list)) Symtab.table * unit Symtab.table | |
| 155 | }; | |
| 156 | ||
| 32544 | 157 | fun make_spec (history_concluded, (eqns, (dtyps, cases))) = | 
| 158 |   Spec { history_concluded = history_concluded, eqns = eqns, dtyps = dtyps, cases = cases };
 | |
| 159 | fun map_spec f (Spec { history_concluded = history_concluded, eqns = eqns,
 | |
| 31962 | 160 | dtyps = dtyps, cases = cases }) = | 
| 32544 | 161 | make_spec (f (history_concluded, (eqns, (dtyps, cases)))); | 
| 162 | fun merge_spec (Spec { history_concluded = _, eqns = eqns1,
 | |
| 31962 | 163 | dtyps = dtyps1, cases = (cases1, undefs1) }, | 
| 32544 | 164 |   Spec { history_concluded = _, eqns = eqns2,
 | 
| 31962 | 165 | dtyps = dtyps2, cases = (cases2, undefs2) }) = | 
| 31156 | 166 | let | 
| 31962 | 167 | fun merge_eqns ((_, history1), (_, history2)) = | 
| 168 | let | |
| 169 | val raw_history = AList.merge (op = : serial * serial -> bool) | |
| 170 | (K true) (history1, history2) | |
| 171 | val filtered_history = filter_out (fst o snd) raw_history | |
| 172 | val history = if null filtered_history | |
| 173 | then raw_history else filtered_history; | |
| 174 | in ((false, (snd o hd) history), history) end; | |
| 175 | val eqns = Symtab.join (K merge_eqns) (eqns1, eqns2); | |
| 176 | val dtyps = Symtab.join (K (AList.merge (op =) (K true))) (dtyps1, dtyps2); | |
| 177 | val cases = (Symtab.merge (K true) (cases1, cases2), | |
| 178 | Symtab.merge (K true) (undefs1, undefs2)); | |
| 32544 | 179 | in make_spec (false, (eqns, (dtyps, cases))) end; | 
| 31962 | 180 | |
| 181 | fun history_concluded (Spec { history_concluded, ... }) = history_concluded;
 | |
| 182 | fun the_eqns (Spec { eqns, ... }) = eqns;
 | |
| 183 | fun the_dtyps (Spec { dtyps, ... }) = dtyps;
 | |
| 184 | fun the_cases (Spec { cases, ... }) = cases;
 | |
| 32544 | 185 | val map_history_concluded = map_spec o apfst; | 
| 31962 | 186 | val map_eqns = map_spec o apsnd o apfst; | 
| 187 | val map_dtyps = map_spec o apsnd o apsnd o apfst; | |
| 188 | val map_cases = map_spec o apsnd o apsnd o apsnd; | |
| 189 | ||
| 190 | ||
| 191 | (* data slots dependent on executable code *) | |
| 192 | ||
| 193 | (*private copy avoids potential conflict of table exceptions*) | |
| 31971 
8c1b845ed105
renamed functor TableFun to Table, and GraphFun to Graph;
 wenzelm parents: 
31962diff
changeset | 194 | structure Datatab = Table(type key = int val ord = int_ord); | 
| 31962 | 195 | |
| 196 | local | |
| 197 | ||
| 198 | type kind = {
 | |
| 199 | empty: Object.T, | |
| 200 | purge: theory -> string list -> Object.T -> Object.T | |
| 201 | }; | |
| 202 | ||
| 32738 | 203 | val kinds = Unsynchronized.ref (Datatab.empty: kind Datatab.table); | 
| 204 | val kind_keys = Unsynchronized.ref ([]: serial list); | |
| 31962 | 205 | |
| 206 | fun invoke f k = case Datatab.lookup (! kinds) k | |
| 207 | of SOME kind => f kind | |
| 208 | | NONE => sys_error "Invalid code data identifier"; | |
| 209 | ||
| 210 | in | |
| 211 | ||
| 212 | fun declare_data empty purge = | |
| 213 | let | |
| 214 | val k = serial (); | |
| 215 |     val kind = {empty = empty, purge = purge};
 | |
| 32738 | 216 | val _ = Unsynchronized.change kinds (Datatab.update (k, kind)); | 
| 217 | val _ = Unsynchronized.change kind_keys (cons k); | |
| 31962 | 218 | in k end; | 
| 219 | ||
| 220 | fun invoke_init k = invoke (fn kind => #empty kind) k; | |
| 221 | ||
| 222 | fun invoke_purge_all thy cs = | |
| 223 | fold (fn k => Datatab.map_entry k | |
| 224 | (invoke (fn kind => #purge kind thy cs) k)) (! kind_keys); | |
| 225 | ||
| 226 | end; (*local*) | |
| 227 | ||
| 228 | ||
| 229 | (* theory store *) | |
| 230 | ||
| 231 | local | |
| 232 | ||
| 233 | type data = Object.T Datatab.table; | |
| 234 | val empty_data = Datatab.empty : data; | |
| 235 | ||
| 236 | structure Code_Data = TheoryDataFun | |
| 237 | ( | |
| 32738 | 238 | type T = spec * data Unsynchronized.ref; | 
| 32544 | 239 | val empty = (make_spec (false, | 
| 32738 | 240 | (Symtab.empty, (Symtab.empty, (Symtab.empty, Symtab.empty)))), Unsynchronized.ref empty_data); | 
| 241 | fun copy (spec, data) = (spec, Unsynchronized.ref (! data)); | |
| 31962 | 242 | val extend = copy; | 
| 243 | fun merge pp ((spec1, data1), (spec2, data2)) = | |
| 32738 | 244 | (merge_spec (spec1, spec2), Unsynchronized.ref empty_data); | 
| 31962 | 245 | ); | 
| 246 | ||
| 247 | fun thy_data f thy = f ((snd o Code_Data.get) thy); | |
| 248 | ||
| 249 | fun get_ensure_init kind data_ref = | |
| 250 | case Datatab.lookup (! data_ref) kind | |
| 251 | of SOME x => x | |
| 252 | | NONE => let val y = invoke_init kind | |
| 32738 | 253 | in (Unsynchronized.change data_ref (Datatab.update (kind, y)); y) end; | 
| 31962 | 254 | |
| 255 | in | |
| 256 | ||
| 257 | (* access to executable code *) | |
| 258 | ||
| 259 | val the_exec = fst o Code_Data.get; | |
| 260 | ||
| 261 | fun complete_class_params thy cs = | |
| 262 | fold (fn c => case AxClass.inst_of_param thy c | |
| 263 | of NONE => insert (op =) c | |
| 264 | | SOME (c', _) => insert (op =) c' #> insert (op =) c) cs []; | |
| 265 | ||
| 266 | fun map_exec_purge touched f thy = | |
| 32738 | 267 | Code_Data.map (fn (exec, data) => (f exec, Unsynchronized.ref (case touched | 
| 31962 | 268 | of SOME cs => invoke_purge_all thy (complete_class_params thy cs) (! data) | 
| 269 | | NONE => empty_data))) thy; | |
| 270 | ||
| 32738 | 271 | val purge_data = (Code_Data.map o apsnd) (K (Unsynchronized.ref empty_data)); | 
| 31962 | 272 | |
| 273 | fun change_eqns delete c f = (map_exec_purge (SOME [c]) o map_eqns | |
| 33006 | 274 | o (if delete then Symtab.map_entry c else Symtab.map_default (c, ((false, (true, [])), []))) | 
| 31962 | 275 | o apfst) (fn (_, eqns) => (true, f eqns)); | 
| 276 | ||
| 33006 | 277 | fun del_eqns c = change_eqns true c (K (false, [])); | 
| 31962 | 278 | |
| 279 | ||
| 280 | (* tackling equation history *) | |
| 281 | ||
| 282 | fun get_eqns thy c = | |
| 283 | Symtab.lookup ((the_eqns o the_exec) thy) c | |
| 33006 | 284 | |> Option.map (snd o snd o fst) | 
| 31962 | 285 | |> these; | 
| 286 | ||
| 287 | fun continue_history thy = if (history_concluded o the_exec) thy | |
| 288 | then thy | |
| 289 | |> (Code_Data.map o apfst o map_history_concluded) (K false) | |
| 290 | |> SOME | |
| 291 | else NONE; | |
| 292 | ||
| 293 | fun conclude_history thy = if (history_concluded o the_exec) thy | |
| 294 | then NONE | |
| 295 | else thy | |
| 296 | |> (Code_Data.map o apfst) | |
| 297 | ((map_eqns o Symtab.map) (fn ((changed, current), history) => | |
| 298 | ((false, current), | |
| 299 | if changed then (serial (), current) :: history else history)) | |
| 300 | #> map_history_concluded (K true)) | |
| 301 | |> SOME; | |
| 302 | ||
| 303 | val _ = Context.>> (Context.map_theory (Code_Data.init | |
| 304 | #> Theory.at_begin continue_history | |
| 305 | #> Theory.at_end conclude_history)); | |
| 306 | ||
| 307 | ||
| 308 | (* access to data dependent on abstract executable code *) | |
| 309 | ||
| 310 | fun get_data (kind, _, dest) = thy_data (get_ensure_init kind #> dest); | |
| 311 | ||
| 312 | fun change_data (kind, mk, dest) = | |
| 313 | let | |
| 314 | fun chnge data_ref f = | |
| 315 | let | |
| 316 | val data = get_ensure_init kind data_ref; | |
| 317 | val data' = f (dest data); | |
| 32738 | 318 | in (Unsynchronized.change data_ref (Datatab.update (kind, mk data')); data') end; | 
| 31962 | 319 | in thy_data chnge end; | 
| 320 | ||
| 321 | fun change_yield_data (kind, mk, dest) = | |
| 322 | let | |
| 323 | fun chnge data_ref f = | |
| 324 | let | |
| 325 | val data = get_ensure_init kind data_ref; | |
| 326 | val (x, data') = f (dest data); | |
| 32738 | 327 | in (x, (Unsynchronized.change data_ref (Datatab.update (kind, mk data')); data')) end; | 
| 31962 | 328 | in thy_data chnge end; | 
| 329 | ||
| 330 | end; (*local*) | |
| 331 | ||
| 332 | ||
| 333 | (** foundation **) | |
| 334 | ||
| 335 | (* constants *) | |
| 336 | ||
| 337 | fun args_number thy = length o fst o strip_type o Sign.the_const_type thy; | |
| 338 | ||
| 339 | ||
| 340 | (* datatypes *) | |
| 31156 | 341 | |
| 342 | fun constrset_of_consts thy cs = | |
| 343 | let | |
| 344 | val _ = map (fn (c, _) => if (is_some o AxClass.class_of_param thy) c | |
| 345 |       then error ("Is a class parameter: " ^ string_of_const thy c) else ()) cs;
 | |
| 33531 | 346 |     fun no_constr s (c, ty) = error ("Not a datatype constructor:\n" ^ string_of_const thy c
 | 
| 347 |       ^ " :: " ^ string_of_typ thy ty ^ "\n" ^ enclose "(" ")" s);
 | |
| 31156 | 348 | fun last_typ c_ty ty = | 
| 349 | let | |
| 33531 | 350 | val tfrees = Term.add_tfreesT ty []; | 
| 31156 | 351 | val (tyco, vs) = ((apsnd o map) (dest_TFree) o dest_Type o snd o strip_type) ty | 
| 33531 | 352 | handle TYPE _ => no_constr "bad type" c_ty | 
| 353 | val _ = if has_duplicates (eq_fst (op =)) vs | |
| 354 | then no_constr "duplicate type variables in datatype" c_ty else (); | |
| 355 | val _ = if length tfrees <> length vs | |
| 356 | then no_constr "type variables missing in datatype" c_ty else (); | |
| 31156 | 357 | in (tyco, vs) end; | 
| 358 | fun ty_sorts (c, ty) = | |
| 359 | let | |
| 360 | val ty_decl = (Logic.unvarifyT o Sign.the_const_type thy) c; | |
| 361 | val (tyco, _) = last_typ (c, ty) ty_decl; | |
| 362 | val (_, vs) = last_typ (c, ty) ty; | |
| 363 | in ((tyco, map snd vs), (c, (map fst vs, ty))) end; | |
| 364 | fun add ((tyco', sorts'), c) ((tyco, sorts), cs) = | |
| 365 | let | |
| 31998 
2c7a24f74db9
code attributes use common underscore convention
 haftmann parents: 
31971diff
changeset | 366 | val _ = if (tyco' : string) <> tyco | 
| 31156 | 367 | then error "Different type constructors in constructor set" | 
| 368 | else (); | |
| 369 | val sorts'' = map2 (curry (Sorts.inter_sort (Sign.classes_of thy))) sorts' sorts | |
| 370 | in ((tyco, sorts), c :: cs) end; | |
| 371 | fun inst vs' (c, (vs, ty)) = | |
| 372 | let | |
| 373 | val the_v = the o AList.lookup (op =) (vs ~~ vs'); | |
| 374 | val ty' = map_atyps (fn TFree (v, _) => TFree (the_v v)) ty; | |
| 375 | in (c, (fst o strip_type) ty') end; | |
| 376 | val c' :: cs' = map ty_sorts cs; | |
| 377 | val ((tyco, sorts), cs'') = fold add cs' (apsnd single c'); | |
| 378 | val vs = Name.names Name.context Name.aT sorts; | |
| 379 | val cs''' = map (inst vs) cs''; | |
| 380 | in (tyco, (vs, rev cs''')) end; | |
| 381 | ||
| 31962 | 382 | fun get_datatype thy tyco = | 
| 383 | case these (Symtab.lookup ((the_dtyps o the_exec) thy) tyco) | |
| 384 | of (_, spec) :: _ => spec | |
| 385 | | [] => Sign.arity_number thy tyco | |
| 386 | |> Name.invents Name.context Name.aT | |
| 387 | |> map (rpair []) | |
| 388 | |> rpair []; | |
| 389 | ||
| 390 | fun get_datatype_of_constr thy c = | |
| 391 | case (snd o strip_type o Sign.the_const_type thy) c | |
| 392 | of Type (tyco, _) => if member (op =) ((map fst o snd o get_datatype thy) tyco) c | |
| 393 | then SOME tyco else NONE | |
| 394 | | _ => NONE; | |
| 395 | ||
| 396 | fun is_constr thy = is_some o get_datatype_of_constr thy; | |
| 397 | ||
| 31156 | 398 | |
| 399 | (* code equations *) | |
| 400 | ||
| 401 | exception BAD_THM of string; | |
| 402 | fun bad_thm msg = raise BAD_THM msg; | |
| 403 | fun error_thm f thm = f thm handle BAD_THM msg => error msg; | |
| 31642 | 404 | fun warning_thm f thm = SOME (f thm) handle BAD_THM msg => (warning msg; NONE) | 
| 31156 | 405 | fun try_thm f thm = SOME (f thm) handle BAD_THM _ => NONE; | 
| 406 | ||
| 407 | fun is_linear thm = | |
| 408 | let val (_, args) = (strip_comb o fst o Logic.dest_equals o Thm.plain_prop_of) thm | |
| 409 | in not (has_duplicates (op =) ((fold o fold_aterms) | |
| 410 | (fn Var (v, _) => cons v | _ => I) args [])) end; | |
| 411 | ||
| 31962 | 412 | fun gen_assert_eqn thy is_constr_pat (thm, proper) = | 
| 31156 | 413 | let | 
| 414 | val (lhs, rhs) = (Logic.dest_equals o Thm.plain_prop_of) thm | |
| 32091 
30e2ffbba718
proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
 wenzelm parents: 
32070diff
changeset | 415 |       handle TERM _ => bad_thm ("Not an equation: " ^ Display.string_of_thm_global thy thm)
 | 
| 
30e2ffbba718
proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
 wenzelm parents: 
32070diff
changeset | 416 |            | THM _ => bad_thm ("Not an equation: " ^ Display.string_of_thm_global thy thm);
 | 
| 31156 | 417 | fun vars_of t = fold_aterms (fn Var (v, _) => insert (op =) v | 
| 418 |       | Free _ => bad_thm ("Illegal free variable in equation\n"
 | |
| 32091 
30e2ffbba718
proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
 wenzelm parents: 
32070diff
changeset | 419 | ^ Display.string_of_thm_global thy thm) | 
| 31156 | 420 | | _ => I) t []; | 
| 421 | fun tvars_of t = fold_term_types (fn _ => | |
| 422 | fold_atyps (fn TVar (v, _) => insert (op =) v | |
| 423 | | TFree _ => bad_thm | |
| 32091 
30e2ffbba718
proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
 wenzelm parents: 
32070diff
changeset | 424 |       ("Illegal free type variable in equation\n" ^ Display.string_of_thm_global thy thm))) t [];
 | 
| 31156 | 425 | val lhs_vs = vars_of lhs; | 
| 426 | val rhs_vs = vars_of rhs; | |
| 427 | val lhs_tvs = tvars_of lhs; | |
| 428 | val rhs_tvs = tvars_of rhs; | |
| 429 | val _ = if null (subtract (op =) lhs_vs rhs_vs) | |
| 430 | then () | |
| 431 |       else bad_thm ("Free variables on right hand side of equation\n"
 | |
| 32091 
30e2ffbba718
proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
 wenzelm parents: 
32070diff
changeset | 432 | ^ Display.string_of_thm_global thy thm); | 
| 31156 | 433 | val _ = if null (subtract (op =) lhs_tvs rhs_tvs) | 
| 434 | then () | |
| 435 |       else bad_thm ("Free type variables on right hand side of equation\n"
 | |
| 32091 
30e2ffbba718
proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
 wenzelm parents: 
32070diff
changeset | 436 | ^ Display.string_of_thm_global thy thm) | 
| 
30e2ffbba718
proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
 wenzelm parents: 
32070diff
changeset | 437 | val (head, args) = (strip_comb o fst o Logic.dest_equals o Thm.plain_prop_of) thm; | 
| 31156 | 438 | val (c, ty) = case head | 
| 439 | of Const (c_ty as (_, ty)) => (AxClass.unoverload_const thy c_ty, ty) | |
| 32091 
30e2ffbba718
proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
 wenzelm parents: 
32070diff
changeset | 440 |       | _ => bad_thm ("Equation not headed by constant\n" ^ Display.string_of_thm_global thy thm);
 | 
| 31156 | 441 | fun check _ (Abs _) = bad_thm | 
| 442 |           ("Abstraction on left hand side of equation\n"
 | |
| 32091 
30e2ffbba718
proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
 wenzelm parents: 
32070diff
changeset | 443 | ^ Display.string_of_thm_global thy thm) | 
| 31156 | 444 | | check 0 (Var _) = () | 
| 445 | | check _ (Var _) = bad_thm | |
| 446 |           ("Variable with application on left hand side of equation\n"
 | |
| 32091 
30e2ffbba718
proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
 wenzelm parents: 
32070diff
changeset | 447 | ^ Display.string_of_thm_global thy thm) | 
| 31156 | 448 | | check n (t1 $ t2) = (check (n+1) t1; check 0 t2) | 
| 449 | | check n (Const (c_ty as (c, ty))) = if n = (length o fst o strip_type) ty | |
| 450 | then if not proper orelse is_constr_pat (AxClass.unoverload_const thy c_ty) | |
| 451 | then () | |
| 452 | else bad_thm (quote c ^ " is not a constructor, on left hand side of equation\n" | |
| 32091 
30e2ffbba718
proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
 wenzelm parents: 
32070diff
changeset | 453 | ^ Display.string_of_thm_global thy thm) | 
| 31156 | 454 | else bad_thm | 
| 455 |             ("Partially applied constant " ^ quote c ^ " on left hand side of equation\n"
 | |
| 32091 
30e2ffbba718
proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
 wenzelm parents: 
32070diff
changeset | 456 | ^ Display.string_of_thm_global thy thm); | 
| 31156 | 457 | val _ = map (check 0) args; | 
| 458 | val _ = if not proper orelse is_linear thm then () | |
| 459 |       else bad_thm ("Duplicate variables on left hand side of equation\n"
 | |
| 32091 
30e2ffbba718
proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
 wenzelm parents: 
32070diff
changeset | 460 | ^ Display.string_of_thm_global thy thm); | 
| 31156 | 461 | val _ = if (is_none o AxClass.class_of_param thy) c | 
| 462 | then () | |
| 463 |       else bad_thm ("Polymorphic constant as head in equation\n"
 | |
| 32091 
30e2ffbba718
proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
 wenzelm parents: 
32070diff
changeset | 464 | ^ Display.string_of_thm_global thy thm) | 
| 31962 | 465 | val _ = if not (is_constr thy c) | 
| 31156 | 466 | then () | 
| 467 |       else bad_thm ("Constructor as head in equation\n"
 | |
| 32091 
30e2ffbba718
proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
 wenzelm parents: 
32070diff
changeset | 468 | ^ Display.string_of_thm_global thy thm) | 
| 31156 | 469 | val ty_decl = Sign.the_const_type thy c; | 
| 470 | val _ = if Sign.typ_equiv thy (Type.strip_sorts ty_decl, Type.strip_sorts ty) | |
| 471 |       then () else bad_thm ("Type\n" ^ string_of_typ thy ty
 | |
| 472 | ^ "\nof equation\n" | |
| 32091 
30e2ffbba718
proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
 wenzelm parents: 
32070diff
changeset | 473 | ^ Display.string_of_thm_global thy thm | 
| 31156 | 474 | ^ "\nis incompatible with declared function type\n" | 
| 475 | ^ string_of_typ thy ty_decl) | |
| 476 | in (thm, proper) end; | |
| 477 | ||
| 31962 | 478 | fun assert_eqn thy = error_thm (gen_assert_eqn thy (is_constr thy)); | 
| 479 | ||
| 480 | fun meta_rewrite thy = LocalDefs.meta_rewrite_rule (ProofContext.init thy); | |
| 31156 | 481 | |
| 31962 | 482 | fun mk_eqn thy = error_thm (gen_assert_eqn thy (K true)) o | 
| 483 | apfst (meta_rewrite thy); | |
| 484 | ||
| 485 | fun mk_eqn_warning thy = Option.map (fn (thm, _) => (thm, is_linear thm)) | |
| 486 | o warning_thm (gen_assert_eqn thy (K true)) o rpair false o meta_rewrite thy; | |
| 487 | ||
| 488 | fun mk_eqn_liberal thy = Option.map (fn (thm, _) => (thm, is_linear thm)) | |
| 489 | o try_thm (gen_assert_eqn thy (K true)) o rpair false o meta_rewrite thy; | |
| 31156 | 490 | |
| 491 | (*those following are permissive wrt. to overloaded constants!*) | |
| 492 | ||
| 32640 
ba6531df2c64
corrected order of type variables in code equations; more precise certificate for cases
 haftmann parents: 
32544diff
changeset | 493 | val head_eqn = dest_Const o fst o strip_comb o fst o Logic.dest_equals o Thm.plain_prop_of; | 
| 31957 | 494 | fun const_typ_eqn thy thm = | 
| 31156 | 495 | let | 
| 32640 
ba6531df2c64
corrected order of type variables in code equations; more precise certificate for cases
 haftmann parents: 
32544diff
changeset | 496 | val (c, ty) = head_eqn thm; | 
| 31156 | 497 | val c' = AxClass.unoverload_const thy (c, ty); | 
| 498 | in (c', ty) end; | |
| 31957 | 499 | fun const_eqn thy = fst o const_typ_eqn thy; | 
| 31156 | 500 | |
| 32873 | 501 | fun typscheme thy (c, ty) = | 
| 502 | (map dest_TFree (Sign.const_typargs thy (c, ty)), Type.strip_sorts ty); | |
| 503 | fun typscheme_eqn thy = typscheme thy o apsnd Logic.unvarifyT o const_typ_eqn thy; | |
| 504 | fun typscheme_eqns thy c [] = | |
| 505 | let | |
| 506 | val raw_ty = Sign.the_const_type thy c; | |
| 507 | val tvars = Term.add_tvar_namesT raw_ty []; | |
| 508 | val tvars' = case AxClass.class_of_param thy c | |
| 509 | of SOME class => [TFree (Name.aT, [class])] | |
| 510 | | NONE => Name.invent_list [] Name.aT (length tvars) | |
| 511 | |> map (fn v => TFree (v, [])); | |
| 512 | val ty = typ_subst_TVars (tvars ~~ tvars') raw_ty; | |
| 513 | in typscheme thy (c, ty) end | |
| 514 | | typscheme_eqns thy c (thms as thm :: _) = typscheme_eqn thy thm; | |
| 32872 
019201eb7e07
variables in type schemes must be renamed simultaneously with variables in equations
 haftmann parents: 
32738diff
changeset | 515 | |
| 31962 | 516 | fun assert_eqns_const thy c eqns = | 
| 517 | let | |
| 518 | fun cert (eqn as (thm, _)) = if c = const_eqn thy thm | |
| 519 |       then eqn else error ("Wrong head of code equation,\nexpected constant "
 | |
| 32091 
30e2ffbba718
proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
 wenzelm parents: 
32070diff
changeset | 520 | ^ string_of_const thy c ^ "\n" ^ Display.string_of_thm_global thy thm) | 
| 31962 | 521 | in map (cert o assert_eqn thy) eqns end; | 
| 31156 | 522 | |
| 32929 | 523 | fun standard_typscheme thy thms = | 
| 32354 | 524 | let | 
| 32640 
ba6531df2c64
corrected order of type variables in code equations; more precise certificate for cases
 haftmann parents: 
32544diff
changeset | 525 | fun tvars_of T = rev (Term.add_tvarsT T []); | 
| 
ba6531df2c64
corrected order of type variables in code equations; more precise certificate for cases
 haftmann parents: 
32544diff
changeset | 526 | val vss = map (tvars_of o snd o head_eqn) thms; | 
| 32354 | 527 | fun inter_sorts vs = | 
| 528 | fold (curry (Sorts.inter_sort (Sign.classes_of thy)) o snd) vs []; | |
| 529 | val sorts = map_transpose inter_sorts vss; | |
| 32873 | 530 | val vts = Name.names Name.context Name.aT sorts | 
| 32354 | 531 | |> map (fn (v, sort) => TVar ((v, 0), sort)); | 
| 532 | in map2 (fn vs => Thm.certify_instantiate (vs ~~ vts, [])) vss thms end; | |
| 31962 | 533 | |
| 534 | fun these_eqns thy c = | |
| 535 | get_eqns thy c | |
| 536 | |> (map o apfst) (Thm.transfer thy) | |
| 32929 | 537 | |> burrow_fst (standard_typscheme thy); | 
| 31962 | 538 | |
| 539 | fun all_eqns thy = | |
| 540 | Symtab.dest ((the_eqns o the_exec) thy) | |
| 33006 | 541 | |> maps (snd o snd o fst o snd); | 
| 31962 | 542 | |
| 543 | ||
| 544 | (* cases *) | |
| 31156 | 545 | |
| 546 | fun case_certificate thm = | |
| 547 | let | |
| 548 | val ((head, raw_case_expr), cases) = (apfst Logic.dest_equals | |
| 32640 
ba6531df2c64
corrected order of type variables in code equations; more precise certificate for cases
 haftmann parents: 
32544diff
changeset | 549 | o apsnd Logic.dest_conjunctions o Logic.dest_implies o Thm.plain_prop_of) thm; | 
| 31156 | 550 | val _ = case head of Free _ => true | 
| 551 | | Var _ => true | |
| 552 |       | _ => raise TERM ("case_cert", []);
 | |
| 553 | val ([(case_var, _)], case_expr) = Term.strip_abs_eta 1 raw_case_expr; | |
| 554 | val (Const (case_const, _), raw_params) = strip_comb case_expr; | |
| 555 | val n = find_index (fn Free (v, _) => v = case_var | _ => false) raw_params; | |
| 556 |     val _ = if n = ~1 then raise TERM ("case_cert", []) else ();
 | |
| 557 | val params = map (fst o dest_Var) (nth_drop n raw_params); | |
| 558 | fun dest_case t = | |
| 559 | let | |
| 560 | val (head' $ t_co, rhs) = Logic.dest_equals t; | |
| 561 |         val _ = if head' = head then () else raise TERM ("case_cert", []);
 | |
| 562 | val (Const (co, _), args) = strip_comb t_co; | |
| 563 | val (Var (param, _), args') = strip_comb rhs; | |
| 564 |         val _ = if args' = args then () else raise TERM ("case_cert", []);
 | |
| 565 | in (param, co) end; | |
| 566 | fun analyze_cases cases = | |
| 567 | let | |
| 568 | val co_list = fold (AList.update (op =) o dest_case) cases []; | |
| 569 | in map (the o AList.lookup (op =) co_list) params end; | |
| 570 | fun analyze_let t = | |
| 571 | let | |
| 572 | val (head' $ arg, Var (param', _) $ arg') = Logic.dest_equals t; | |
| 573 |         val _ = if head' = head then () else raise TERM ("case_cert", []);
 | |
| 574 |         val _ = if arg' = arg then () else raise TERM ("case_cert", []);
 | |
| 575 |         val _ = if [param'] = params then () else raise TERM ("case_cert", []);
 | |
| 576 | in [] end; | |
| 577 | fun analyze (cases as [let_case]) = | |
| 578 | (analyze_cases cases handle Bind => analyze_let let_case) | |
| 579 | | analyze cases = analyze_cases cases; | |
| 580 | in (case_const, (n, analyze cases)) end; | |
| 581 | ||
| 582 | fun case_cert thm = case_certificate thm | |
| 583 | handle Bind => error "bad case certificate" | |
| 584 | | TERM _ => error "bad case certificate"; | |
| 585 | ||
| 31962 | 586 | fun get_case_scheme thy = Symtab.lookup ((fst o the_cases o the_exec) thy); | 
| 24219 | 587 | |
| 31962 | 588 | val undefineds = Symtab.keys o snd o the_cases o the_exec; | 
| 24219 | 589 | |
| 590 | ||
| 31962 | 591 | (* diagnostic *) | 
| 24219 | 592 | |
| 593 | fun print_codesetup thy = | |
| 594 | let | |
| 595 | val ctxt = ProofContext.init thy; | |
| 24844 
98c006a30218
certificates for code generator case expressions
 haftmann parents: 
24837diff
changeset | 596 | val exec = the_exec thy; | 
| 33006 | 597 | fun pretty_eqns (s, (_, eqns)) = | 
| 24219 | 598 | (Pretty.block o Pretty.fbreaks) ( | 
| 33006 | 599 | Pretty.str s :: map (Display.pretty_thm ctxt o fst) eqns | 
| 24219 | 600 | ); | 
| 601 | fun pretty_dtyp (s, []) = | |
| 602 | Pretty.str s | |
| 603 | | pretty_dtyp (s, cos) = | |
| 604 | (Pretty.block o Pretty.breaks) ( | |
| 605 | Pretty.str s | |
| 606 | :: Pretty.str "=" | |
| 31156 | 607 | :: separate (Pretty.str "|") (map (fn (c, []) => Pretty.str (string_of_const thy c) | 
| 24219 | 608 | | (c, tys) => | 
| 609 | (Pretty.block o Pretty.breaks) | |
| 31156 | 610 | (Pretty.str (string_of_const thy c) | 
| 26947 
133905a0c493
moved global pretty/string_of functions from Sign to Syntax;
 wenzelm parents: 
26928diff
changeset | 611 | :: Pretty.str "of" | 
| 
133905a0c493
moved global pretty/string_of functions from Sign to Syntax;
 wenzelm parents: 
26928diff
changeset | 612 | :: map (Pretty.quote o Syntax.pretty_typ_global thy) tys)) cos) | 
| 24219 | 613 | ); | 
| 28368 | 614 | val eqns = the_eqns exec | 
| 24423 
ae9cd0e92423
overloaded definitions accompanied by explicit constants
 haftmann parents: 
24283diff
changeset | 615 | |> Symtab.dest | 
| 31156 | 616 | |> (map o apfst) (string_of_const thy) | 
| 28695 | 617 | |> (map o apsnd) (snd o fst) | 
| 24219 | 618 | |> sort (string_ord o pairself fst); | 
| 619 | val dtyps = the_dtyps exec | |
| 620 | |> Symtab.dest | |
| 28695 | 621 | |> map (fn (dtco, (_, (vs, cos)) :: _) => | 
| 31957 | 622 | (string_of_typ thy (Type (dtco, map TFree vs)), cos)) | 
| 24219 | 623 | |> sort (string_ord o pairself fst) | 
| 624 | in | |
| 625 | (Pretty.writeln o Pretty.chunks) [ | |
| 626 | Pretty.block ( | |
| 30023 
55954f726043
permissive check for pattern discipline in case schemes
 haftmann parents: 
29970diff
changeset | 627 | Pretty.str "code equations:" | 
| 24219 | 628 | :: Pretty.fbrk | 
| 33006 | 629 | :: (Pretty.fbreaks o map pretty_eqns) eqns | 
| 24219 | 630 | ), | 
| 25968 | 631 | Pretty.block ( | 
| 24219 | 632 | Pretty.str "datatypes:" | 
| 633 | :: Pretty.fbrk | |
| 634 | :: (Pretty.fbreaks o map pretty_dtyp) dtyps | |
| 635 | ) | |
| 636 | ] | |
| 637 | end; | |
| 638 | ||
| 639 | ||
| 31962 | 640 | (** declaring executable ingredients **) | 
| 641 | ||
| 642 | (* datatypes *) | |
| 31090 
3be41b271023
clarified matter of "proper" flag in code equations
 haftmann parents: 
31088diff
changeset | 643 | |
| 33314 | 644 | structure Type_Interpretation = | 
| 645 | Interpretation(type T = string * serial val eq = eq_snd (op =) : T * T -> bool); | |
| 25462 | 646 | |
| 24423 
ae9cd0e92423
overloaded definitions accompanied by explicit constants
 haftmann parents: 
24283diff
changeset | 647 | fun add_datatype raw_cs thy = | 
| 24219 | 648 | let | 
| 25597 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 haftmann parents: 
25501diff
changeset | 649 | val cs = map (fn c_ty as (_, ty) => (AxClass.unoverload_const thy c_ty, ty)) raw_cs; | 
| 31156 | 650 | val (tyco, vs_cos) = constrset_of_consts thy cs; | 
| 30076 | 651 | val old_cs = (map fst o snd o get_datatype thy) tyco; | 
| 652 | fun drop_outdated_cases cases = fold Symtab.delete_safe | |
| 653 | (Symtab.fold (fn (c, (_, (_, cos))) => | |
| 654 | if exists (member (op =) old_cs) cos | |
| 655 | then insert (op =) c else I) cases []) cases; | |
| 24219 | 656 | in | 
| 657 | thy | |
| 31091 
8316d22f10f5
corrected deletetion of code equations for constructors
 haftmann parents: 
31090diff
changeset | 658 | |> fold (del_eqns o fst) cs | 
| 28703 | 659 | |> map_exec_purge NONE | 
| 28695 | 660 | ((map_dtyps o Symtab.map_default (tyco, [])) (cons (serial (), vs_cos)) | 
| 30076 | 661 | #> (map_cases o apfst) drop_outdated_cases) | 
| 31962 | 662 | |> Type_Interpretation.data (tyco, serial ()) | 
| 24219 | 663 | end; | 
| 664 | ||
| 31962 | 665 | fun type_interpretation f = Type_Interpretation.interpretation | 
| 25485 | 666 | (fn (tyco, _) => fn thy => f (tyco, get_datatype thy tyco) thy); | 
| 667 | ||
| 24423 
ae9cd0e92423
overloaded definitions accompanied by explicit constants
 haftmann parents: 
24283diff
changeset | 668 | fun add_datatype_cmd raw_cs thy = | 
| 
ae9cd0e92423
overloaded definitions accompanied by explicit constants
 haftmann parents: 
24283diff
changeset | 669 | let | 
| 31156 | 670 | val cs = map (read_bare_const thy) raw_cs; | 
| 24423 
ae9cd0e92423
overloaded definitions accompanied by explicit constants
 haftmann parents: 
24283diff
changeset | 671 | in add_datatype cs thy end; | 
| 24219 | 672 | |
| 31962 | 673 | |
| 674 | (* code equations *) | |
| 675 | ||
| 33075 | 676 | fun gen_add_eqn default (thm, proper) thy = | 
| 677 | let | |
| 678 | val thm' = Thm.close_derivation thm; | |
| 679 | val c = const_eqn thy thm'; | |
| 680 | in change_eqns false c (add_thm thy default (thm', proper)) thy end; | |
| 31962 | 681 | |
| 682 | fun add_eqn thm thy = | |
| 683 | gen_add_eqn false (mk_eqn thy (thm, true)) thy; | |
| 684 | ||
| 685 | fun add_warning_eqn thm thy = | |
| 686 | case mk_eqn_warning thy thm | |
| 687 | of SOME eqn => gen_add_eqn false eqn thy | |
| 688 | | NONE => thy; | |
| 689 | ||
| 690 | fun add_default_eqn thm thy = | |
| 691 | case mk_eqn_liberal thy thm | |
| 692 | of SOME eqn => gen_add_eqn true eqn thy | |
| 693 | | NONE => thy; | |
| 694 | ||
| 695 | fun add_nbe_eqn thm thy = | |
| 696 | gen_add_eqn false (mk_eqn thy (thm, false)) thy; | |
| 697 | ||
| 698 | val add_default_eqn_attribute = Thm.declaration_attribute | |
| 699 | (fn thm => Context.mapping (add_default_eqn thm) I); | |
| 700 | val add_default_eqn_attrib = Attrib.internal (K add_default_eqn_attribute); | |
| 701 | ||
| 702 | fun del_eqn thm thy = case mk_eqn_liberal thy thm | |
| 703 | of SOME (thm, _) => change_eqns true (const_eqn thy thm) (del_thm thm) thy | |
| 704 | | NONE => thy; | |
| 705 | ||
| 32070 | 706 | (* c.f. src/HOL/Tools/recfun_codegen.ML *) | 
| 707 | ||
| 33522 | 708 | structure Code_Target_Attr = Theory_Data | 
| 709 | ( | |
| 31998 
2c7a24f74db9
code attributes use common underscore convention
 haftmann parents: 
31971diff
changeset | 710 | type T = (string -> thm -> theory -> theory) option; | 
| 
2c7a24f74db9
code attributes use common underscore convention
 haftmann parents: 
31971diff
changeset | 711 | val empty = NONE; | 
| 
2c7a24f74db9
code attributes use common underscore convention
 haftmann parents: 
31971diff
changeset | 712 | val extend = I; | 
| 33522 | 713 | fun merge (f1, f2) = if is_some f1 then f1 else f2; | 
| 31998 
2c7a24f74db9
code attributes use common underscore convention
 haftmann parents: 
31971diff
changeset | 714 | ); | 
| 
2c7a24f74db9
code attributes use common underscore convention
 haftmann parents: 
31971diff
changeset | 715 | |
| 32070 | 716 | fun set_code_target_attr f = Code_Target_Attr.map (K (SOME f)); | 
| 717 | ||
| 718 | fun code_target_attr prefix thm thy = | |
| 719 | let | |
| 720 | val attr = the_default ((K o K) I) (Code_Target_Attr.get thy); | |
| 721 | in thy |> add_warning_eqn thm |> attr prefix thm end; | |
| 722 | ||
| 723 | (* setup *) | |
| 31998 
2c7a24f74db9
code attributes use common underscore convention
 haftmann parents: 
31971diff
changeset | 724 | |
| 31962 | 725 | val _ = Context.>> (Context.map_theory | 
| 726 | (let | |
| 727 | fun mk_attribute f = Thm.declaration_attribute (fn thm => Context.mapping (f thm) I); | |
| 31998 
2c7a24f74db9
code attributes use common underscore convention
 haftmann parents: 
31971diff
changeset | 728 | val code_attribute_parser = | 
| 
2c7a24f74db9
code attributes use common underscore convention
 haftmann parents: 
31971diff
changeset | 729 | Args.del |-- Scan.succeed (mk_attribute del_eqn) | 
| 
2c7a24f74db9
code attributes use common underscore convention
 haftmann parents: 
31971diff
changeset | 730 | || Args.$$$ "nbe" |-- Scan.succeed (mk_attribute add_nbe_eqn) | 
| 
2c7a24f74db9
code attributes use common underscore convention
 haftmann parents: 
31971diff
changeset | 731 | || (Args.$$$ "target" |-- Args.colon |-- Args.name >> | 
| 32070 | 732 | (mk_attribute o code_target_attr)) | 
| 31998 
2c7a24f74db9
code attributes use common underscore convention
 haftmann parents: 
31971diff
changeset | 733 | || Scan.succeed (mk_attribute add_warning_eqn); | 
| 31962 | 734 | in | 
| 735 | Type_Interpretation.init | |
| 31998 
2c7a24f74db9
code attributes use common underscore convention
 haftmann parents: 
31971diff
changeset | 736 | #> Attrib.setup (Binding.name "code") (Scan.lift code_attribute_parser) | 
| 
2c7a24f74db9
code attributes use common underscore convention
 haftmann parents: 
31971diff
changeset | 737 | "declare theorems for code generation" | 
| 31962 | 738 | end)); | 
| 739 | ||
| 740 | ||
| 741 | (* cases *) | |
| 742 | ||
| 24844 
98c006a30218
certificates for code generator case expressions
 haftmann parents: 
24837diff
changeset | 743 | fun add_case thm thy = | 
| 
98c006a30218
certificates for code generator case expressions
 haftmann parents: 
24837diff
changeset | 744 | let | 
| 31156 | 745 | val (c, (k, case_pats)) = case_cert thm; | 
| 31088 | 746 | val _ = case filter_out (is_constr thy) case_pats | 
| 30076 | 747 | of [] => () | 
| 748 |       | cs => error ("Non-constructor(s) in case certificate: " ^ commas (map quote cs));
 | |
| 749 | val entry = (1 + Int.max (1, length case_pats), (k, case_pats)) | |
| 750 | in (map_exec_purge (SOME [c]) o map_cases o apfst) (Symtab.update (c, entry)) thy end; | |
| 24844 
98c006a30218
certificates for code generator case expressions
 haftmann parents: 
24837diff
changeset | 751 | |
| 
98c006a30218
certificates for code generator case expressions
 haftmann parents: 
24837diff
changeset | 752 | fun add_undefined c thy = | 
| 
98c006a30218
certificates for code generator case expressions
 haftmann parents: 
24837diff
changeset | 753 | (map_exec_purge (SOME [c]) o map_cases o apsnd) (Symtab.update (c, ())) thy; | 
| 
98c006a30218
certificates for code generator case expressions
 haftmann parents: 
24837diff
changeset | 754 | |
| 24219 | 755 | end; (*struct*) | 
| 756 | ||
| 757 | ||
| 32661 
f4d179d91af2
experimenting to add some useful interface for definitions
 bulwahn parents: 
32640diff
changeset | 758 | (** type-safe interfaces for data dependent on executable code **) | 
| 24219 | 759 | |
| 31962 | 760 | functor Code_Data_Fun(Data: CODE_DATA_ARGS): CODE_DATA = | 
| 24219 | 761 | struct | 
| 762 | ||
| 763 | type T = Data.T; | |
| 764 | exception Data of T; | |
| 765 | fun dest (Data x) = x | |
| 766 | ||
| 767 | val kind = Code.declare_data (Data Data.empty) | |
| 27609 | 768 | (fn thy => fn cs => fn Data x => Data (Data.purge thy cs x)); | 
| 24219 | 769 | |
| 770 | val data_op = (kind, Data, dest); | |
| 771 | ||
| 772 | val get = Code.get_data data_op; | |
| 773 | val change = Code.change_data data_op; | |
| 774 | fun change_yield thy = Code.change_yield_data data_op thy; | |
| 775 | ||
| 776 | end; | |
| 777 | ||
| 32661 
f4d179d91af2
experimenting to add some useful interface for definitions
 bulwahn parents: 
32640diff
changeset | 778 | (** datastructure to log definitions for preprocessing of the predicate compiler **) | 
| 32662 | 779 | (* obviously a clone of Named_Thms *) | 
| 780 | ||
| 781 | signature PREDICATE_COMPILE_PREPROC_CONST_DEFS = | |
| 782 | sig | |
| 783 | val get: Proof.context -> thm list | |
| 784 | val add_thm: thm -> Context.generic -> Context.generic | |
| 785 | val del_thm: thm -> Context.generic -> Context.generic | |
| 786 | ||
| 787 | val add_attribute : attribute | |
| 788 | val del_attribute : attribute | |
| 789 | ||
| 790 | val add_attrib : Attrib.src | |
| 791 | ||
| 792 | val setup: theory -> theory | |
| 793 | end; | |
| 794 | ||
| 795 | structure Predicate_Compile_Preproc_Const_Defs : PREDICATE_COMPILE_PREPROC_CONST_DEFS = | |
| 796 | struct | |
| 797 | ||
| 33519 | 798 | structure Data = Generic_Data | 
| 32661 
f4d179d91af2
experimenting to add some useful interface for definitions
 bulwahn parents: 
32640diff
changeset | 799 | ( | 
| 32662 | 800 | type T = thm list; | 
| 801 | val empty = []; | |
| 802 | val extend = I; | |
| 33519 | 803 | val merge = Thm.merge_thms; | 
| 32662 | 804 | ); | 
| 805 | ||
| 806 | val get = Data.get o Context.Proof; | |
| 807 | ||
| 808 | val add_thm = Data.map o Thm.add_thm; | |
| 809 | val del_thm = Data.map o Thm.del_thm; | |
| 810 | ||
| 811 | val add_attribute = Thm.declaration_attribute add_thm; | |
| 812 | val del_attribute = Thm.declaration_attribute del_thm; | |
| 813 | ||
| 814 | val add_attrib = Attrib.internal (K add_attribute) | |
| 815 | ||
| 816 | val setup = | |
| 817 | Attrib.setup (Binding.name "pred_compile_preproc") (Attrib.add_del add_attribute del_attribute) | |
| 818 |     ("declaration of definition for preprocessing of the predicate compiler") #>
 | |
| 819 | PureThy.add_thms_dynamic (Binding.name "pred_compile_preproc", Data.get); | |
| 820 | ||
| 821 | end; | |
| 822 | ||
| 28143 | 823 | structure Code : CODE = struct open Code; end; |