| author | boehmes | 
| Wed, 24 Nov 2010 08:51:48 +0100 | |
| changeset 40680 | 02caa72cb950 | 
| parent 40564 | 6827505e96e1 | 
| child 40726 | 16dcfedc4eb7 | 
| permissions | -rw-r--r-- | 
| 24219 | 1 | (* Title: Pure/Isar/code.ML | 
| 2 | Author: Florian Haftmann, TU Muenchen | |
| 3 | ||
| 34173 
458ced35abb8
reduced code generator cache to the baremost minimum
 haftmann parents: 
33977diff
changeset | 4 | Abstract executable ingredients of theory. Management of data | 
| 
458ced35abb8
reduced code generator cache to the baremost minimum
 haftmann parents: 
33977diff
changeset | 5 | dependent on executable ingredients as synchronized cache; purged | 
| 
458ced35abb8
reduced code generator cache to the baremost minimum
 haftmann parents: 
33977diff
changeset | 6 | on any change of underlying executable ingredients. | 
| 24219 | 7 | *) | 
| 8 | ||
| 9 | signature CODE = | |
| 10 | sig | |
| 31957 | 11 | (*constants*) | 
| 12 | val check_const: theory -> term -> string | |
| 13 | val read_bare_const: theory -> string -> string * typ | |
| 14 | val read_const: theory -> string -> string | |
| 31962 | 15 | val string_of_const: theory -> string -> string | 
| 33940 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 16 | val cert_signature: theory -> typ -> typ | 
| 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 17 | val read_signature: theory -> string -> typ | 
| 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 18 | val const_typ: theory -> string -> typ | 
| 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 19 | val subst_signatures: theory -> term -> term | 
| 31962 | 20 | val args_number: theory -> string -> int | 
| 31957 | 21 | |
| 31156 | 22 | (*constructor sets*) | 
| 23 | val constrset_of_consts: theory -> (string * typ) list | |
| 24 | -> string * ((string * sort) list * (string * typ list) list) | |
| 25 | ||
| 34874 | 26 | (*code equations and certificates*) | 
| 31962 | 27 | val mk_eqn: theory -> thm * bool -> thm * bool | 
| 28 | val mk_eqn_warning: theory -> thm -> (thm * bool) option | |
| 29 | val mk_eqn_liberal: theory -> thm -> (thm * bool) option | |
| 31156 | 30 | val assert_eqn: theory -> thm * bool -> thm * bool | 
| 31957 | 31 | val const_typ_eqn: theory -> thm -> string * typ | 
| 34895 | 32 | val expand_eta: theory -> int -> thm -> thm | 
| 33 | type cert | |
| 34891 
99b9a6290446
code certificates as integral part of code generation
 haftmann parents: 
34874diff
changeset | 34 | val empty_cert: theory -> string -> cert | 
| 
99b9a6290446
code certificates as integral part of code generation
 haftmann parents: 
34874diff
changeset | 35 | val cert_of_eqns: theory -> string -> (thm * bool) list -> cert | 
| 34874 | 36 | val constrain_cert: theory -> sort list -> cert -> cert | 
| 35226 | 37 | val typargs_deps_of_cert: theory -> cert -> (string * sort) list * (string * typ list) list | 
| 36209 
566be5d48090
more appropriate representation of valid positions for abstractors; more accurate checking of abstype certificates
 haftmann parents: 
36202diff
changeset | 38 | val equations_of_cert: theory -> cert -> ((string * sort) list * typ) | 
| 
566be5d48090
more appropriate representation of valid positions for abstractors; more accurate checking of abstype certificates
 haftmann parents: 
36202diff
changeset | 39 | * (((term * string option) list * (term * string option)) * (thm option * bool)) list | 
| 35226 | 40 | val bare_thms_of_cert: theory -> cert -> thm list | 
| 34895 | 41 | val pretty_cert: theory -> cert -> Pretty.T list | 
| 31156 | 42 | |
| 31962 | 43 | (*executable code*) | 
| 33940 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 44 | val add_type: string -> theory -> theory | 
| 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 45 | val add_type_cmd: string -> theory -> theory | 
| 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 46 | val add_signature: string * typ -> theory -> theory | 
| 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 47 | val add_signature_cmd: string * string -> theory -> theory | 
| 31156 | 48 | val add_datatype: (string * typ) list -> theory -> theory | 
| 49 | val add_datatype_cmd: string list -> theory -> theory | |
| 35299 
4f4d5bf4ea08
proper distinction of code datatypes and abstypes
 haftmann parents: 
35226diff
changeset | 50 | val datatype_interpretation: | 
| 
4f4d5bf4ea08
proper distinction of code datatypes and abstypes
 haftmann parents: 
35226diff
changeset | 51 | (string * ((string * sort) list * (string * typ list) list) | 
| 
4f4d5bf4ea08
proper distinction of code datatypes and abstypes
 haftmann parents: 
35226diff
changeset | 52 | -> theory -> theory) -> theory -> theory | 
| 36112 
7fa17a225852
user interface for abstract datatypes is an attribute, not a command
 haftmann parents: 
35845diff
changeset | 53 | val add_abstype: thm -> theory -> theory | 
| 35299 
4f4d5bf4ea08
proper distinction of code datatypes and abstypes
 haftmann parents: 
35226diff
changeset | 54 | val abstype_interpretation: | 
| 
4f4d5bf4ea08
proper distinction of code datatypes and abstypes
 haftmann parents: 
35226diff
changeset | 55 | (string * ((string * sort) list * ((string * typ) * (string * thm))) | 
| 31156 | 56 | -> theory -> theory) -> theory -> theory | 
| 28368 | 57 | val add_eqn: thm -> theory -> theory | 
| 31088 | 58 | val add_nbe_eqn: thm -> theory -> theory | 
| 28368 | 59 | val add_default_eqn: thm -> theory -> theory | 
| 28703 | 60 | val add_default_eqn_attribute: attribute | 
| 61 | val add_default_eqn_attrib: Attrib.src | |
| 37425 | 62 | val add_nbe_default_eqn: thm -> theory -> theory | 
| 63 | val add_nbe_default_eqn_attribute: attribute | |
| 64 | val add_nbe_default_eqn_attrib: Attrib.src | |
| 28368 | 65 | val del_eqn: thm -> theory -> theory | 
| 66 | val del_eqns: string -> theory -> theory | |
| 24844 
98c006a30218
certificates for code generator case expressions
 haftmann parents: 
24837diff
changeset | 67 | val add_case: thm -> theory -> theory | 
| 
98c006a30218
certificates for code generator case expressions
 haftmann parents: 
24837diff
changeset | 68 | val add_undefined: string -> theory -> theory | 
| 37448 
3bd4b3809bee
explicit type variable arguments for constructors
 haftmann parents: 
37438diff
changeset | 69 | val get_type: theory -> string -> ((string * sort) list * ((string * string list) * typ list) list) | 
| 35299 
4f4d5bf4ea08
proper distinction of code datatypes and abstypes
 haftmann parents: 
35226diff
changeset | 70 | val get_type_of_constr_or_abstr: theory -> string -> (string * bool) option | 
| 35226 | 71 | val is_constr: theory -> string -> bool | 
| 72 | val is_abstr: theory -> string -> bool | |
| 34891 
99b9a6290446
code certificates as integral part of code generation
 haftmann parents: 
34874diff
changeset | 73 | val get_cert: theory -> ((thm * bool) list -> (thm * bool) list) -> string -> cert | 
| 30023 
55954f726043
permissive check for pattern discipline in case schemes
 haftmann parents: 
29970diff
changeset | 74 | val get_case_scheme: theory -> string -> (int * (int * string list)) option | 
| 37438 | 75 | val get_case_cong: theory -> string -> thm option | 
| 31890 
e943b039f0ac
an intermediate step towards a refined translation of cases
 haftmann parents: 
31642diff
changeset | 76 | val undefineds: theory -> string list | 
| 24219 | 77 | val print_codesetup: theory -> unit | 
| 31957 | 78 | |
| 79 | (*infrastructure*) | |
| 31998 
2c7a24f74db9
code attributes use common underscore convention
 haftmann parents: 
31971diff
changeset | 80 | val set_code_target_attr: (string -> thm -> theory -> theory) -> theory -> theory | 
| 24219 | 81 | end; | 
| 82 | ||
| 83 | signature CODE_DATA_ARGS = | |
| 84 | sig | |
| 85 | type T | |
| 86 | val empty: T | |
| 87 | end; | |
| 88 | ||
| 89 | signature CODE_DATA = | |
| 90 | sig | |
| 91 | type T | |
| 39397 | 92 | val change: theory option -> (T -> T) -> T | 
| 93 | val change_yield: theory option -> (T -> 'a * T) -> 'a * T | |
| 24219 | 94 | end; | 
| 95 | ||
| 96 | signature PRIVATE_CODE = | |
| 97 | sig | |
| 98 | include CODE | |
| 34173 
458ced35abb8
reduced code generator cache to the baremost minimum
 haftmann parents: 
33977diff
changeset | 99 | val declare_data: Object.T -> serial | 
| 24219 | 100 |   val change_yield_data: serial * ('a -> Object.T) * (Object.T -> 'a)
 | 
| 34251 | 101 |     -> theory -> ('a -> 'b * 'a) -> 'b * 'a
 | 
| 24219 | 102 | end; | 
| 103 | ||
| 104 | structure Code : PRIVATE_CODE = | |
| 105 | struct | |
| 106 | ||
| 31962 | 107 | (** auxiliary **) | 
| 108 | ||
| 109 | (* printing *) | |
| 31156 | 110 | |
| 39134 
917b4b6ba3d2
turned show_sorts/show_types into proper configuration options;
 wenzelm parents: 
39020diff
changeset | 111 | fun string_of_typ thy = | 
| 
917b4b6ba3d2
turned show_sorts/show_types into proper configuration options;
 wenzelm parents: 
39020diff
changeset | 112 | Syntax.string_of_typ (Config.put show_sorts true (Syntax.init_pretty_global thy)); | 
| 31962 | 113 | |
| 31156 | 114 | fun string_of_const thy c = case AxClass.inst_of_param thy c | 
| 115 | of SOME (c, tyco) => Sign.extern_const thy c ^ " " ^ enclose "[" "]" (Sign.extern_type thy tyco) | |
| 116 | | NONE => Sign.extern_const thy c; | |
| 117 | ||
| 31962 | 118 | |
| 119 | (* constants *) | |
| 31156 | 120 | |
| 33940 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 121 | fun typ_equiv tys = Type.raw_instance tys andalso Type.raw_instance (swap tys); | 
| 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 122 | |
| 31962 | 123 | fun check_bare_const thy t = case try dest_Const t | 
| 124 | of SOME c_ty => c_ty | |
| 125 |   | NONE => error ("Not a constant: " ^ Syntax.string_of_term_global thy t);
 | |
| 31156 | 126 | |
| 40362 
82a066bff182
Code.check_const etc.: reject too specific types
 haftmann parents: 
40316diff
changeset | 127 | fun check_unoverload thy (c, ty) = | 
| 
82a066bff182
Code.check_const etc.: reject too specific types
 haftmann parents: 
40316diff
changeset | 128 | let | 
| 
82a066bff182
Code.check_const etc.: reject too specific types
 haftmann parents: 
40316diff
changeset | 129 | val c' = AxClass.unoverload_const thy (c, ty); | 
| 
82a066bff182
Code.check_const etc.: reject too specific types
 haftmann parents: 
40316diff
changeset | 130 | val ty_decl = Sign.the_const_type thy c'; | 
| 
82a066bff182
Code.check_const etc.: reject too specific types
 haftmann parents: 
40316diff
changeset | 131 | in if Sign.typ_equiv thy | 
| 
82a066bff182
Code.check_const etc.: reject too specific types
 haftmann parents: 
40316diff
changeset | 132 | (Type.strip_sorts ty_decl, Type.strip_sorts (Logic.varifyT_global ty)) then c' | 
| 
82a066bff182
Code.check_const etc.: reject too specific types
 haftmann parents: 
40316diff
changeset | 133 |     else error ("Type\n" ^ string_of_typ thy ty
 | 
| 
82a066bff182
Code.check_const etc.: reject too specific types
 haftmann parents: 
40316diff
changeset | 134 | ^ "\nof constant " ^ quote c | 
| 
82a066bff182
Code.check_const etc.: reject too specific types
 haftmann parents: 
40316diff
changeset | 135 | ^ "\nis too specific compared to declared type\n" | 
| 
82a066bff182
Code.check_const etc.: reject too specific types
 haftmann parents: 
40316diff
changeset | 136 | ^ string_of_typ thy ty_decl) | 
| 
82a066bff182
Code.check_const etc.: reject too specific types
 haftmann parents: 
40316diff
changeset | 137 | end; | 
| 
82a066bff182
Code.check_const etc.: reject too specific types
 haftmann parents: 
40316diff
changeset | 138 | |
| 
82a066bff182
Code.check_const etc.: reject too specific types
 haftmann parents: 
40316diff
changeset | 139 | fun check_const thy = check_unoverload thy o check_bare_const thy; | 
| 31962 | 140 | |
| 141 | fun read_bare_const thy = check_bare_const thy o Syntax.read_term_global thy; | |
| 142 | ||
| 40362 
82a066bff182
Code.check_const etc.: reject too specific types
 haftmann parents: 
40316diff
changeset | 143 | fun read_const thy = check_unoverload thy o read_bare_const thy; | 
| 31156 | 144 | |
| 32872 
019201eb7e07
variables in type schemes must be renamed simultaneously with variables in equations
 haftmann parents: 
32738diff
changeset | 145 | |
| 31962 | 146 | (** data store **) | 
| 147 | ||
| 35226 | 148 | (* datatypes *) | 
| 149 | ||
| 150 | datatype typ_spec = Constructors of (string * typ list) list | |
| 151 | | Abstractor of (string * typ) * (string * thm); | |
| 31962 | 152 | |
| 35226 | 153 | fun constructors_of (Constructors cos) = (cos, false) | 
| 154 | | constructors_of (Abstractor ((co, ty), _)) = ([(co, [ty])], true); | |
| 155 | ||
| 156 | ||
| 157 | (* functions *) | |
| 31962 | 158 | |
| 37460 
910b2422571d
drop subsumed default equations (requires a little bit unfortunate laziness)
 haftmann parents: 
37448diff
changeset | 159 | datatype fun_spec = Default of (thm * bool) list * (thm * bool) list lazy | 
| 35226 | 160 | | Eqns of (thm * bool) list | 
| 161 | | Proj of term * string | |
| 162 | | Abstr of thm * string; | |
| 31962 | 163 | |
| 37460 
910b2422571d
drop subsumed default equations (requires a little bit unfortunate laziness)
 haftmann parents: 
37448diff
changeset | 164 | val empty_fun_spec = Default ([], Lazy.value []); | 
| 31962 | 165 | |
| 35226 | 166 | fun is_default (Default _) = true | 
| 167 | | is_default _ = false; | |
| 168 | ||
| 169 | fun associated_abstype (Abstr (_, tyco)) = SOME tyco | |
| 170 | | associated_abstype _ = NONE; | |
| 31962 | 171 | |
| 172 | ||
| 173 | (* executable code data *) | |
| 174 | ||
| 175 | datatype spec = Spec of {
 | |
| 176 | history_concluded: bool, | |
| 33940 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 177 | signatures: int Symtab.table * typ Symtab.table, | 
| 35226 | 178 | functions: ((bool * fun_spec) * (serial * fun_spec) list) Symtab.table | 
| 31962 | 179 | (*with explicit history*), | 
| 35299 
4f4d5bf4ea08
proper distinction of code datatypes and abstypes
 haftmann parents: 
35226diff
changeset | 180 | types: ((serial * ((string * sort) list * typ_spec)) list) Symtab.table | 
| 31962 | 181 | (*with explicit history*), | 
| 37438 | 182 | cases: ((int * (int * string list)) * thm) Symtab.table * unit Symtab.table | 
| 31962 | 183 | }; | 
| 184 | ||
| 35299 
4f4d5bf4ea08
proper distinction of code datatypes and abstypes
 haftmann parents: 
35226diff
changeset | 185 | fun make_spec (history_concluded, ((signatures, functions), (types, cases))) = | 
| 33940 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 186 |   Spec { history_concluded = history_concluded,
 | 
| 35299 
4f4d5bf4ea08
proper distinction of code datatypes and abstypes
 haftmann parents: 
35226diff
changeset | 187 | signatures = signatures, functions = functions, types = types, cases = cases }; | 
| 33940 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 188 | fun map_spec f (Spec { history_concluded = history_concluded, signatures = signatures,
 | 
| 35299 
4f4d5bf4ea08
proper distinction of code datatypes and abstypes
 haftmann parents: 
35226diff
changeset | 189 | functions = functions, types = types, cases = cases }) = | 
| 
4f4d5bf4ea08
proper distinction of code datatypes and abstypes
 haftmann parents: 
35226diff
changeset | 190 | make_spec (f (history_concluded, ((signatures, functions), (types, cases)))); | 
| 35226 | 191 | fun merge_spec (Spec { history_concluded = _, signatures = (tycos1, sigs1), functions = functions1,
 | 
| 35299 
4f4d5bf4ea08
proper distinction of code datatypes and abstypes
 haftmann parents: 
35226diff
changeset | 192 | types = types1, cases = (cases1, undefs1) }, | 
| 35226 | 193 |   Spec { history_concluded = _, signatures = (tycos2, sigs2), functions = functions2,
 | 
| 35299 
4f4d5bf4ea08
proper distinction of code datatypes and abstypes
 haftmann parents: 
35226diff
changeset | 194 | types = types2, cases = (cases2, undefs2) }) = | 
| 31156 | 195 | let | 
| 33940 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 196 | val signatures = (Symtab.merge (op =) (tycos1, tycos2), | 
| 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 197 | Symtab.merge typ_equiv (sigs1, sigs2)); | 
| 35226 | 198 | fun merge_functions ((_, history1), (_, history2)) = | 
| 31962 | 199 | let | 
| 200 | val raw_history = AList.merge (op = : serial * serial -> bool) | |
| 35226 | 201 | (K true) (history1, history2); | 
| 202 | val filtered_history = filter_out (is_default o snd) raw_history; | |
| 31962 | 203 | val history = if null filtered_history | 
| 204 | then raw_history else filtered_history; | |
| 205 | in ((false, (snd o hd) history), history) end; | |
| 35226 | 206 | val functions = Symtab.join (K merge_functions) (functions1, functions2); | 
| 35299 
4f4d5bf4ea08
proper distinction of code datatypes and abstypes
 haftmann parents: 
35226diff
changeset | 207 | val types = Symtab.join (K (AList.merge (op =) (K true))) (types1, types2); | 
| 31962 | 208 | val cases = (Symtab.merge (K true) (cases1, cases2), | 
| 209 | Symtab.merge (K true) (undefs1, undefs2)); | |
| 35299 
4f4d5bf4ea08
proper distinction of code datatypes and abstypes
 haftmann parents: 
35226diff
changeset | 210 | in make_spec (false, ((signatures, functions), (types, cases))) end; | 
| 31962 | 211 | |
| 212 | fun history_concluded (Spec { history_concluded, ... }) = history_concluded;
 | |
| 33940 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 213 | fun the_signatures (Spec { signatures, ... }) = signatures;
 | 
| 35226 | 214 | fun the_functions (Spec { functions, ... }) = functions;
 | 
| 35299 
4f4d5bf4ea08
proper distinction of code datatypes and abstypes
 haftmann parents: 
35226diff
changeset | 215 | fun the_types (Spec { types, ... }) = types;
 | 
| 31962 | 216 | fun the_cases (Spec { cases, ... }) = cases;
 | 
| 32544 | 217 | val map_history_concluded = map_spec o apfst; | 
| 33940 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 218 | val map_signatures = map_spec o apsnd o apfst o apfst; | 
| 35226 | 219 | val map_functions = map_spec o apsnd o apfst o apsnd; | 
| 220 | val map_typs = map_spec o apsnd o apsnd o apfst; | |
| 31962 | 221 | val map_cases = map_spec o apsnd o apsnd o apsnd; | 
| 222 | ||
| 223 | ||
| 224 | (* data slots dependent on executable code *) | |
| 225 | ||
| 226 | (*private copy avoids potential conflict of table exceptions*) | |
| 31971 
8c1b845ed105
renamed functor TableFun to Table, and GraphFun to Graph;
 wenzelm parents: 
31962diff
changeset | 227 | structure Datatab = Table(type key = int val ord = int_ord); | 
| 31962 | 228 | |
| 229 | local | |
| 230 | ||
| 34173 
458ced35abb8
reduced code generator cache to the baremost minimum
 haftmann parents: 
33977diff
changeset | 231 | type kind = { empty: Object.T };
 | 
| 31962 | 232 | |
| 32738 | 233 | val kinds = Unsynchronized.ref (Datatab.empty: kind Datatab.table); | 
| 31962 | 234 | |
| 235 | fun invoke f k = case Datatab.lookup (! kinds) k | |
| 236 | of SOME kind => f kind | |
| 40316 
665862241968
replaced ancient sys_error by raise Fail, assuming that the latter is not handled specifically by the environment;
 wenzelm parents: 
40187diff
changeset | 237 | | NONE => raise Fail "Invalid code data identifier"; | 
| 31962 | 238 | |
| 239 | in | |
| 240 | ||
| 34173 
458ced35abb8
reduced code generator cache to the baremost minimum
 haftmann parents: 
33977diff
changeset | 241 | fun declare_data empty = | 
| 31962 | 242 | let | 
| 243 | val k = serial (); | |
| 34173 
458ced35abb8
reduced code generator cache to the baremost minimum
 haftmann parents: 
33977diff
changeset | 244 |     val kind = { empty = empty };
 | 
| 
458ced35abb8
reduced code generator cache to the baremost minimum
 haftmann parents: 
33977diff
changeset | 245 | val _ = CRITICAL (fn () => Unsynchronized.change kinds (Datatab.update (k, kind))); | 
| 31962 | 246 | in k end; | 
| 247 | ||
| 248 | fun invoke_init k = invoke (fn kind => #empty kind) k; | |
| 249 | ||
| 250 | end; (*local*) | |
| 251 | ||
| 252 | ||
| 253 | (* theory store *) | |
| 254 | ||
| 255 | local | |
| 256 | ||
| 257 | type data = Object.T Datatab.table; | |
| 34244 | 258 | fun empty_dataref () = Synchronized.var "code data" (NONE : (data * theory_ref) option); | 
| 31962 | 259 | |
| 34244 | 260 | structure Code_Data = Theory_Data | 
| 31962 | 261 | ( | 
| 34244 | 262 | type T = spec * (data * theory_ref) option Synchronized.var; | 
| 33940 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 263 | val empty = (make_spec (false, (((Symtab.empty, Symtab.empty), Symtab.empty), | 
| 34244 | 264 | (Symtab.empty, (Symtab.empty, Symtab.empty)))), empty_dataref ()); | 
| 265 | val extend = I | |
| 266 | fun merge ((spec1, _), (spec2, _)) = | |
| 267 | (merge_spec (spec1, spec2), empty_dataref ()); | |
| 31962 | 268 | ); | 
| 269 | ||
| 270 | in | |
| 271 | ||
| 35226 | 272 | |
| 31962 | 273 | (* access to executable code *) | 
| 274 | ||
| 275 | val the_exec = fst o Code_Data.get; | |
| 276 | ||
| 34244 | 277 | fun map_exec_purge f = Code_Data.map (fn (exec, _) => (f exec, empty_dataref ())); | 
| 31962 | 278 | |
| 35226 | 279 | fun change_fun_spec delete c f = (map_exec_purge o map_functions | 
| 280 | o (if delete then Symtab.map_entry c else Symtab.map_default (c, ((false, empty_fun_spec), []))) | |
| 281 | o apfst) (fn (_, spec) => (true, f spec)); | |
| 31962 | 282 | |
| 283 | ||
| 284 | (* tackling equation history *) | |
| 285 | ||
| 286 | fun continue_history thy = if (history_concluded o the_exec) thy | |
| 287 | then thy | |
| 288 | |> (Code_Data.map o apfst o map_history_concluded) (K false) | |
| 289 | |> SOME | |
| 290 | else NONE; | |
| 291 | ||
| 292 | fun conclude_history thy = if (history_concluded o the_exec) thy | |
| 293 | then NONE | |
| 294 | else thy | |
| 295 | |> (Code_Data.map o apfst) | |
| 39020 | 296 | ((map_functions o Symtab.map) (fn _ => fn ((changed, current), history) => | 
| 31962 | 297 | ((false, current), | 
| 298 | if changed then (serial (), current) :: history else history)) | |
| 299 | #> map_history_concluded (K true)) | |
| 300 | |> SOME; | |
| 301 | ||
| 34244 | 302 | val _ = Context.>> (Context.map_theory (Theory.at_begin continue_history #> Theory.at_end conclude_history)); | 
| 31962 | 303 | |
| 304 | ||
| 305 | (* access to data dependent on abstract executable code *) | |
| 306 | ||
| 34244 | 307 | fun change_yield_data (kind, mk, dest) theory f = | 
| 31962 | 308 | let | 
| 34244 | 309 | val dataref = (snd o Code_Data.get) theory; | 
| 34251 | 310 | val (datatab, thy_ref) = case Synchronized.value dataref | 
| 311 | of SOME (datatab, thy_ref) => if Theory.eq_thy (theory, Theory.deref thy_ref) | |
| 312 | then (datatab, thy_ref) | |
| 313 | else (Datatab.empty, Theory.check_thy theory) | |
| 314 | | NONE => (Datatab.empty, Theory.check_thy theory) | |
| 34244 | 315 | val data = case Datatab.lookup datatab kind | 
| 316 | of SOME data => data | |
| 317 | | NONE => invoke_init kind; | |
| 34251 | 318 | val result as (x, data') = f (dest data); | 
| 34244 | 319 | val _ = Synchronized.change dataref | 
| 320 | ((K o SOME) (Datatab.update (kind, mk data') datatab, thy_ref)); | |
| 321 | in result end; | |
| 31962 | 322 | |
| 323 | end; (*local*) | |
| 324 | ||
| 325 | ||
| 326 | (** foundation **) | |
| 327 | ||
| 328 | (* constants *) | |
| 329 | ||
| 33940 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 330 | fun arity_number thy tyco = case Symtab.lookup ((fst o the_signatures o the_exec) thy) tyco | 
| 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 331 | of SOME n => n | 
| 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 332 | | NONE => Sign.arity_number thy tyco; | 
| 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 333 | |
| 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 334 | fun build_tsig thy = | 
| 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 335 | let | 
| 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 336 | val (tycos, _) = (the_signatures o the_exec) thy; | 
| 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 337 | val decls = (#types o Type.rep_tsig o Sign.tsig_of) thy | 
| 34272 | 338 | |> snd | 
| 339 | |> Symtab.fold (fn (tyco, n) => | |
| 340 | Symtab.update (tyco, Type.LogicalType n)) tycos; | |
| 341 | in | |
| 342 | Type.empty_tsig | |
| 343 | |> Symtab.fold (fn (tyco, Type.LogicalType n) => Type.add_type Name_Space.default_naming | |
| 344 | (Binding.qualified_name tyco, n) | _ => I) decls | |
| 345 | end; | |
| 33940 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 346 | |
| 35845 
e5980f0ad025
renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
 wenzelm parents: 
35624diff
changeset | 347 | fun cert_signature thy = Logic.varifyT_global o Type.cert_typ (build_tsig thy) o Type.no_tvars; | 
| 33940 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 348 | |
| 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 349 | fun read_signature thy = cert_signature thy o Type.strip_sorts | 
| 36610 
bafd82950e24
renamed ProofContext.init to ProofContext.init_global to emphasize that this is not the real thing;
 wenzelm parents: 
36209diff
changeset | 350 | o Syntax.parse_typ (ProofContext.init_global thy); | 
| 33940 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 351 | |
| 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 352 | fun expand_signature thy = Type.cert_typ_mode Type.mode_syntax (Sign.tsig_of thy); | 
| 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 353 | |
| 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 354 | fun lookup_typ thy = Symtab.lookup ((snd o the_signatures o the_exec) thy); | 
| 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 355 | |
| 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 356 | fun const_typ thy c = case lookup_typ thy c | 
| 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 357 | of SOME ty => ty | 
| 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 358 | | NONE => (Type.strip_sorts o Sign.the_const_type thy) c; | 
| 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 359 | |
| 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 360 | fun subst_signature thy c ty = | 
| 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 361 | let | 
| 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 362 | fun mk_subst (Type (tyco, tys1)) (ty2 as Type (tyco2, tys2)) = | 
| 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 363 | fold2 mk_subst tys1 tys2 | 
| 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 364 | | mk_subst ty (TVar (v, sort)) = Vartab.update (v, ([], ty)) | 
| 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 365 | in case lookup_typ thy c | 
| 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 366 | of SOME ty' => Envir.subst_type (mk_subst ty (expand_signature thy ty') Vartab.empty) ty' | 
| 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 367 | | NONE => ty | 
| 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 368 | end; | 
| 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 369 | |
| 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 370 | fun subst_signatures thy = map_aterms (fn Const (c, ty) => Const (c, subst_signature thy c ty) | t => t); | 
| 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 371 | |
| 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 372 | fun args_number thy = length o fst o strip_type o const_typ thy; | 
| 31962 | 373 | |
| 374 | ||
| 375 | (* datatypes *) | |
| 31156 | 376 | |
| 35226 | 377 | fun no_constr thy s (c, ty) = error ("Not a datatype constructor:\n" ^ string_of_const thy c
 | 
| 378 |   ^ " :: " ^ string_of_typ thy ty ^ "\n" ^ enclose "(" ")" s);
 | |
| 379 | ||
| 380 | fun ty_sorts thy (c, raw_ty) = | |
| 31156 | 381 | let | 
| 35226 | 382 | val _ = Thm.cterm_of thy (Const (c, raw_ty)); | 
| 383 | val ty = subst_signature thy c raw_ty; | |
| 35845 
e5980f0ad025
renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
 wenzelm parents: 
35624diff
changeset | 384 | val ty_decl = (Logic.unvarifyT_global o const_typ thy) c; | 
| 31156 | 385 | fun last_typ c_ty ty = | 
| 386 | let | |
| 33531 | 387 | val tfrees = Term.add_tfreesT ty []; | 
| 31156 | 388 | val (tyco, vs) = ((apsnd o map) (dest_TFree) o dest_Type o snd o strip_type) ty | 
| 35226 | 389 | handle TYPE _ => no_constr thy "bad type" c_ty | 
| 36112 
7fa17a225852
user interface for abstract datatypes is an attribute, not a command
 haftmann parents: 
35845diff
changeset | 390 | val _ = if tyco = "fun" then no_constr thy "bad type" c_ty else (); | 
| 33531 | 391 | val _ = if has_duplicates (eq_fst (op =)) vs | 
| 35226 | 392 | then no_constr thy "duplicate type variables in datatype" c_ty else (); | 
| 33531 | 393 | val _ = if length tfrees <> length vs | 
| 35226 | 394 | then no_constr thy "type variables missing in datatype" c_ty else (); | 
| 31156 | 395 | in (tyco, vs) end; | 
| 35226 | 396 | val (tyco, _) = last_typ (c, ty) ty_decl; | 
| 397 | val (_, vs) = last_typ (c, ty) ty; | |
| 398 | in ((tyco, map snd vs), (c, (map fst vs, ty))) end; | |
| 399 | ||
| 400 | fun constrset_of_consts thy cs = | |
| 401 | let | |
| 402 | val _ = map (fn (c, _) => if (is_some o AxClass.class_of_param thy) c | |
| 403 |       then error ("Is a class parameter: " ^ string_of_const thy c) else ()) cs;
 | |
| 31156 | 404 | fun add ((tyco', sorts'), c) ((tyco, sorts), cs) = | 
| 405 | let | |
| 31998 
2c7a24f74db9
code attributes use common underscore convention
 haftmann parents: 
31971diff
changeset | 406 | val _ = if (tyco' : string) <> tyco | 
| 31156 | 407 | then error "Different type constructors in constructor set" | 
| 408 | else (); | |
| 409 | val sorts'' = map2 (curry (Sorts.inter_sort (Sign.classes_of thy))) sorts' sorts | |
| 410 | in ((tyco, sorts), c :: cs) end; | |
| 411 | fun inst vs' (c, (vs, ty)) = | |
| 412 | let | |
| 413 | val the_v = the o AList.lookup (op =) (vs ~~ vs'); | |
| 414 | val ty' = map_atyps (fn TFree (v, _) => TFree (the_v v)) ty; | |
| 415 | in (c, (fst o strip_type) ty') end; | |
| 35226 | 416 | val c' :: cs' = map (ty_sorts thy) cs; | 
| 31156 | 417 | val ((tyco, sorts), cs'') = fold add cs' (apsnd single c'); | 
| 418 | val vs = Name.names Name.context Name.aT sorts; | |
| 419 | val cs''' = map (inst vs) cs''; | |
| 420 | in (tyco, (vs, rev cs''')) end; | |
| 421 | ||
| 35299 
4f4d5bf4ea08
proper distinction of code datatypes and abstypes
 haftmann parents: 
35226diff
changeset | 422 | fun get_type_entry thy tyco = case these (Symtab.lookup ((the_types o the_exec) thy) tyco) | 
| 35226 | 423 | of (_, entry) :: _ => SOME entry | 
| 424 | | _ => NONE; | |
| 31962 | 425 | |
| 35299 
4f4d5bf4ea08
proper distinction of code datatypes and abstypes
 haftmann parents: 
35226diff
changeset | 426 | fun get_type_spec thy tyco = case get_type_entry thy tyco | 
| 35226 | 427 | of SOME (vs, spec) => apfst (pair vs) (constructors_of spec) | 
| 428 | | NONE => arity_number thy tyco | |
| 429 | |> Name.invents Name.context Name.aT | |
| 430 | |> map (rpair []) | |
| 431 | |> rpair [] | |
| 432 | |> rpair false; | |
| 433 | ||
| 35299 
4f4d5bf4ea08
proper distinction of code datatypes and abstypes
 haftmann parents: 
35226diff
changeset | 434 | fun get_abstype_spec thy tyco = case get_type_entry thy tyco | 
| 35226 | 435 | of SOME (vs, Abstractor spec) => (vs, spec) | 
| 36122 | 436 |   | _ => error ("Not an abstract type: " ^ tyco);
 | 
| 35226 | 437 | |
| 37448 
3bd4b3809bee
explicit type variable arguments for constructors
 haftmann parents: 
37438diff
changeset | 438 | fun get_type thy tyco = | 
| 
3bd4b3809bee
explicit type variable arguments for constructors
 haftmann parents: 
37438diff
changeset | 439 | let | 
| 
3bd4b3809bee
explicit type variable arguments for constructors
 haftmann parents: 
37438diff
changeset | 440 | val ((vs, cos), _) = get_type_spec thy tyco; | 
| 
3bd4b3809bee
explicit type variable arguments for constructors
 haftmann parents: 
37438diff
changeset | 441 | fun args_of c tys = map (fst o dest_TFree) | 
| 
3bd4b3809bee
explicit type variable arguments for constructors
 haftmann parents: 
37438diff
changeset | 442 | (Sign.const_typargs thy (c, tys ---> Type (tyco, map TFree vs))); | 
| 
3bd4b3809bee
explicit type variable arguments for constructors
 haftmann parents: 
37438diff
changeset | 443 | fun add_typargs (c, tys) = ((c, args_of c tys), tys); | 
| 
3bd4b3809bee
explicit type variable arguments for constructors
 haftmann parents: 
37438diff
changeset | 444 | in (vs, map add_typargs cos) end; | 
| 35226 | 445 | |
| 35299 
4f4d5bf4ea08
proper distinction of code datatypes and abstypes
 haftmann parents: 
35226diff
changeset | 446 | fun get_type_of_constr_or_abstr thy c = | 
| 33940 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 447 | case (snd o strip_type o const_typ thy) c | 
| 35299 
4f4d5bf4ea08
proper distinction of code datatypes and abstypes
 haftmann parents: 
35226diff
changeset | 448 | of Type (tyco, _) => let val ((vs, cos), abstract) = get_type_spec thy tyco | 
| 35226 | 449 | in if member (op =) (map fst cos) c then SOME (tyco, abstract) else NONE end | 
| 31962 | 450 | | _ => NONE; | 
| 451 | ||
| 35299 
4f4d5bf4ea08
proper distinction of code datatypes and abstypes
 haftmann parents: 
35226diff
changeset | 452 | fun is_constr thy c = case get_type_of_constr_or_abstr thy c | 
| 35226 | 453 | of SOME (_, false) => true | 
| 454 | | _ => false; | |
| 455 | ||
| 35299 
4f4d5bf4ea08
proper distinction of code datatypes and abstypes
 haftmann parents: 
35226diff
changeset | 456 | fun is_abstr thy c = case get_type_of_constr_or_abstr thy c | 
| 35226 | 457 | of SOME (_, true) => true | 
| 458 | | _ => false; | |
| 31962 | 459 | |
| 31156 | 460 | |
| 34874 | 461 | (* bare code equations *) | 
| 31156 | 462 | |
| 35226 | 463 | (* convention for variables: | 
| 464 | ?x ?'a for free-floating theorems (e.g. in the data store) | |
| 465 | ?x 'a for certificates | |
| 466 | x 'a for final representation of equations | |
| 467 | *) | |
| 468 | ||
| 31156 | 469 | exception BAD_THM of string; | 
| 470 | fun bad_thm msg = raise BAD_THM msg; | |
| 471 | fun error_thm f thm = f thm handle BAD_THM msg => error msg; | |
| 31642 | 472 | fun warning_thm f thm = SOME (f thm) handle BAD_THM msg => (warning msg; NONE) | 
| 31156 | 473 | fun try_thm f thm = SOME (f thm) handle BAD_THM _ => NONE; | 
| 474 | ||
| 475 | fun is_linear thm = | |
| 476 | let val (_, args) = (strip_comb o fst o Logic.dest_equals o Thm.plain_prop_of) thm | |
| 477 | in not (has_duplicates (op =) ((fold o fold_aterms) | |
| 478 | (fn Var (v, _) => cons v | _ => I) args [])) end; | |
| 479 | ||
| 36209 
566be5d48090
more appropriate representation of valid positions for abstractors; more accurate checking of abstype certificates
 haftmann parents: 
36202diff
changeset | 480 | fun check_decl_ty thy (c, ty) = | 
| 
566be5d48090
more appropriate representation of valid positions for abstractors; more accurate checking of abstype certificates
 haftmann parents: 
36202diff
changeset | 481 | let | 
| 
566be5d48090
more appropriate representation of valid positions for abstractors; more accurate checking of abstype certificates
 haftmann parents: 
36202diff
changeset | 482 | val ty_decl = Sign.the_const_type thy c; | 
| 
566be5d48090
more appropriate representation of valid positions for abstractors; more accurate checking of abstype certificates
 haftmann parents: 
36202diff
changeset | 483 | in if Sign.typ_equiv thy (Type.strip_sorts ty_decl, Type.strip_sorts ty) then () | 
| 
566be5d48090
more appropriate representation of valid positions for abstractors; more accurate checking of abstype certificates
 haftmann parents: 
36202diff
changeset | 484 |     else bad_thm ("Type\n" ^ string_of_typ thy ty
 | 
| 
566be5d48090
more appropriate representation of valid positions for abstractors; more accurate checking of abstype certificates
 haftmann parents: 
36202diff
changeset | 485 | ^ "\nof constant " ^ quote c | 
| 40362 
82a066bff182
Code.check_const etc.: reject too specific types
 haftmann parents: 
40316diff
changeset | 486 | ^ "\nis too specific compared to declared type\n" | 
| 36209 
566be5d48090
more appropriate representation of valid positions for abstractors; more accurate checking of abstype certificates
 haftmann parents: 
36202diff
changeset | 487 | ^ string_of_typ thy ty_decl) | 
| 
566be5d48090
more appropriate representation of valid positions for abstractors; more accurate checking of abstype certificates
 haftmann parents: 
36202diff
changeset | 488 | end; | 
| 
566be5d48090
more appropriate representation of valid positions for abstractors; more accurate checking of abstype certificates
 haftmann parents: 
36202diff
changeset | 489 | |
| 35226 | 490 | fun check_eqn thy { allow_nonlinear, allow_consts, allow_pats } thm (lhs, rhs) =
 | 
| 31156 | 491 | let | 
| 34894 | 492 | fun bad s = bad_thm (s ^ ":\n" ^ Display.string_of_thm_global thy thm); | 
| 31156 | 493 | fun vars_of t = fold_aterms (fn Var (v, _) => insert (op =) v | 
| 34894 | 494 | | Free _ => bad "Illegal free variable in equation" | 
| 31156 | 495 | | _ => I) t []; | 
| 496 | fun tvars_of t = fold_term_types (fn _ => | |
| 497 | fold_atyps (fn TVar (v, _) => insert (op =) v | |
| 34894 | 498 | | TFree _ => bad "Illegal free type variable in equation")) t []; | 
| 31156 | 499 | val lhs_vs = vars_of lhs; | 
| 500 | val rhs_vs = vars_of rhs; | |
| 501 | val lhs_tvs = tvars_of lhs; | |
| 502 | val rhs_tvs = tvars_of rhs; | |
| 503 | val _ = if null (subtract (op =) lhs_vs rhs_vs) | |
| 504 | then () | |
| 34894 | 505 | else bad "Free variables on right hand side of equation"; | 
| 31156 | 506 | val _ = if null (subtract (op =) lhs_tvs rhs_tvs) | 
| 507 | then () | |
| 34894 | 508 | else bad "Free type variables on right hand side of equation"; | 
| 509 | val (head, args) = strip_comb lhs; | |
| 31156 | 510 | val (c, ty) = case head | 
| 511 | of Const (c_ty as (_, ty)) => (AxClass.unoverload_const thy c_ty, ty) | |
| 34894 | 512 | | _ => bad "Equation not headed by constant"; | 
| 513 | fun check _ (Abs _) = bad "Abstraction on left hand side of equation" | |
| 31156 | 514 | | check 0 (Var _) = () | 
| 34894 | 515 | | check _ (Var _) = bad "Variable with application on left hand side of equation" | 
| 31156 | 516 | | check n (t1 $ t2) = (check (n+1) t1; check 0 t2) | 
| 34894 | 517 | | check n (Const (c_ty as (c, ty))) = | 
| 35226 | 518 | if allow_pats then let | 
| 33940 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 519 | val c' = AxClass.unoverload_const thy c_ty | 
| 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 520 | in if n = (length o fst o strip_type o subst_signature thy c') ty | 
| 35226 | 521 | then if allow_consts orelse is_constr thy c' | 
| 33940 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 522 | then () | 
| 34894 | 523 | else bad (quote c ^ " is not a constructor, on left hand side of equation") | 
| 524 |             else bad ("Partially applied constant " ^ quote c ^ " on left hand side of equation")
 | |
| 36209 
566be5d48090
more appropriate representation of valid positions for abstractors; more accurate checking of abstype certificates
 haftmann parents: 
36202diff
changeset | 525 |           end else bad ("Pattern not allowed here, but constant " ^ quote c ^ " encountered on left hand side")
 | 
| 31156 | 526 | val _ = map (check 0) args; | 
| 35226 | 527 | val _ = if allow_nonlinear orelse is_linear thm then () | 
| 34894 | 528 | else bad "Duplicate variables on left hand side of equation"; | 
| 529 | val _ = if (is_none o AxClass.class_of_param thy) c then () | |
| 530 | else bad "Overloaded constant as head in equation"; | |
| 531 | val _ = if not (is_constr thy c) then () | |
| 532 | else bad "Constructor as head in equation"; | |
| 35226 | 533 | val _ = if not (is_abstr thy c) then () | 
| 534 | else bad "Abstractor as head in equation"; | |
| 36209 
566be5d48090
more appropriate representation of valid positions for abstractors; more accurate checking of abstype certificates
 haftmann parents: 
36202diff
changeset | 535 | val _ = check_decl_ty thy (c, ty); | 
| 35226 | 536 | in () end; | 
| 537 | ||
| 538 | fun gen_assert_eqn thy check_patterns (thm, proper) = | |
| 539 | let | |
| 540 | fun bad s = bad_thm (s ^ ":\n" ^ Display.string_of_thm_global thy thm); | |
| 541 | val (lhs, rhs) = (Logic.dest_equals o Thm.plain_prop_of) thm | |
| 542 | handle TERM _ => bad "Not an equation" | |
| 543 | | THM _ => bad "Not a proper equation"; | |
| 544 |     val _ = check_eqn thy { allow_nonlinear = not proper,
 | |
| 545 | allow_consts = not (proper andalso check_patterns), allow_pats = true } thm (lhs, rhs); | |
| 31156 | 546 | in (thm, proper) end; | 
| 547 | ||
| 35226 | 548 | fun assert_abs_eqn thy some_tyco thm = | 
| 549 | let | |
| 550 | fun bad s = bad_thm (s ^ ":\n" ^ Display.string_of_thm_global thy thm); | |
| 551 | val (full_lhs, rhs) = (Logic.dest_equals o Thm.plain_prop_of) thm | |
| 552 | handle TERM _ => bad "Not an equation" | |
| 553 | | THM _ => bad "Not a proper equation"; | |
| 554 | val (rep, lhs) = dest_comb full_lhs | |
| 555 | handle TERM _ => bad "Not an abstract equation"; | |
| 36202 
43ea1f28fc7c
explicit check sorts in abstract certificates; abstractor is part of dependencies
 haftmann parents: 
36122diff
changeset | 556 | val (rep_const, ty) = dest_Const rep; | 
| 40187 
9b6e918682d5
more general treatment of type argument in code certificates for operations on abstract types
 haftmann parents: 
39811diff
changeset | 557 | val (tyco, Ts) = (dest_Type o domain_type) ty | 
| 39552 | 558 | handle TERM _ => bad "Not an abstract equation" | 
| 559 | | TYPE _ => bad "Not an abstract equation"; | |
| 35226 | 560 | val _ = case some_tyco of SOME tyco' => if tyco = tyco' then () | 
| 561 |           else bad ("Abstract type mismatch:" ^ quote tyco ^ " vs. " ^ quote tyco')
 | |
| 562 | | NONE => (); | |
| 36202 
43ea1f28fc7c
explicit check sorts in abstract certificates; abstractor is part of dependencies
 haftmann parents: 
36122diff
changeset | 563 | val (vs', (_, (rep', _))) = get_abstype_spec thy tyco; | 
| 35226 | 564 | val _ = if rep_const = rep' then () | 
| 565 |       else bad ("Projection mismatch: " ^ quote rep_const ^ " vs. " ^ quote rep');
 | |
| 566 |     val _ = check_eqn thy { allow_nonlinear = false,
 | |
| 567 | allow_consts = false, allow_pats = false } thm (lhs, rhs); | |
| 40564 | 568 | val _ = if forall2 (fn T => fn (_, sort) => Sign.of_sort thy (T, sort)) Ts vs' then () | 
| 40187 
9b6e918682d5
more general treatment of type argument in code certificates for operations on abstract types
 haftmann parents: 
39811diff
changeset | 569 |       else error ("Type arguments do not satisfy sort constraints of abstype certificate.");
 | 
| 35226 | 570 | in (thm, tyco) end; | 
| 571 | ||
| 34894 | 572 | fun assert_eqn thy = error_thm (gen_assert_eqn thy true); | 
| 31962 | 573 | |
| 36610 
bafd82950e24
renamed ProofContext.init to ProofContext.init_global to emphasize that this is not the real thing;
 wenzelm parents: 
36209diff
changeset | 574 | fun meta_rewrite thy = Local_Defs.meta_rewrite_rule (ProofContext.init_global thy); | 
| 31156 | 575 | |
| 34894 | 576 | fun mk_eqn thy = error_thm (gen_assert_eqn thy false) o | 
| 31962 | 577 | apfst (meta_rewrite thy); | 
| 578 | ||
| 579 | fun mk_eqn_warning thy = Option.map (fn (thm, _) => (thm, is_linear thm)) | |
| 34894 | 580 | o warning_thm (gen_assert_eqn thy false) o rpair false o meta_rewrite thy; | 
| 31962 | 581 | |
| 582 | fun mk_eqn_liberal thy = Option.map (fn (thm, _) => (thm, is_linear thm)) | |
| 34894 | 583 | o try_thm (gen_assert_eqn thy false) o rpair false o meta_rewrite thy; | 
| 31156 | 584 | |
| 35226 | 585 | fun mk_abs_eqn thy = error_thm (assert_abs_eqn thy NONE) o meta_rewrite thy; | 
| 586 | ||
| 33940 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 587 | val head_eqn = dest_Const o fst o strip_comb o fst o Logic.dest_equals o Thm.plain_prop_of; | 
| 31156 | 588 | |
| 31957 | 589 | fun const_typ_eqn thy thm = | 
| 31156 | 590 | let | 
| 32640 
ba6531df2c64
corrected order of type variables in code equations; more precise certificate for cases
 haftmann parents: 
32544diff
changeset | 591 | val (c, ty) = head_eqn thm; | 
| 31156 | 592 | val c' = AxClass.unoverload_const thy (c, ty); | 
| 33940 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 593 | (*permissive wrt. to overloaded constants!*) | 
| 31156 | 594 | in (c', ty) end; | 
| 33940 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 595 | |
| 31957 | 596 | fun const_eqn thy = fst o const_typ_eqn thy; | 
| 31156 | 597 | |
| 35226 | 598 | fun const_abs_eqn thy = AxClass.unoverload_const thy o dest_Const o fst o strip_comb o snd | 
| 599 | o dest_comb o fst o Logic.dest_equals o Thm.plain_prop_of; | |
| 600 | ||
| 34874 | 601 | fun logical_typscheme thy (c, ty) = | 
| 32873 | 602 | (map dest_TFree (Sign.const_typargs thy (c, ty)), Type.strip_sorts ty); | 
| 33940 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 603 | |
| 34874 | 604 | fun typscheme thy (c, ty) = logical_typscheme thy (c, subst_signature thy c ty); | 
| 33940 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 605 | |
| 35226 | 606 | fun mk_proj tyco vs ty abs rep = | 
| 607 | let | |
| 608 | val ty_abs = Type (tyco, map TFree vs); | |
| 609 |     val xarg = Var (("x", 0), ty);
 | |
| 610 | in Logic.mk_equals (Const (rep, ty_abs --> ty) $ (Const (abs, ty --> ty_abs) $ xarg), xarg) end; | |
| 611 | ||
| 34895 | 612 | |
| 613 | (* technical transformations of code equations *) | |
| 614 | ||
| 615 | fun expand_eta thy k thm = | |
| 616 | let | |
| 617 | val (lhs, rhs) = (Logic.dest_equals o Thm.plain_prop_of) thm; | |
| 618 | val (_, args) = strip_comb lhs; | |
| 619 | val l = if k = ~1 | |
| 620 | then (length o fst o strip_abs) rhs | |
| 621 | else Int.max (0, k - length args); | |
| 622 | val (raw_vars, _) = Term.strip_abs_eta l rhs; | |
| 623 | val vars = burrow_fst (Name.variant_list (map (fst o fst) (Term.add_vars lhs []))) | |
| 624 | raw_vars; | |
| 625 | fun expand (v, ty) thm = Drule.fun_cong_rule thm | |
| 626 | (Thm.cterm_of thy (Var ((v, 0), ty))); | |
| 627 | in | |
| 628 | thm | |
| 629 | |> fold expand vars | |
| 630 | |> Conv.fconv_rule Drule.beta_eta_conversion | |
| 631 | end; | |
| 632 | ||
| 633 | fun same_arity thy thms = | |
| 31962 | 634 | let | 
| 34895 | 635 | val num_args_of = length o snd o strip_comb o fst o Logic.dest_equals; | 
| 636 | val k = fold (Integer.max o num_args_of o Thm.prop_of) thms 0; | |
| 637 | in map (expand_eta thy k) thms end; | |
| 638 | ||
| 639 | fun mk_desymbolization pre post mk vs = | |
| 640 | let | |
| 641 | val names = map (pre o fst o fst) vs | |
| 642 | |> map (Name.desymbolize false) | |
| 643 | |> Name.variant_list [] | |
| 644 | |> map post; | |
| 645 | in map_filter (fn (((v, i), x), v') => | |
| 646 | if v = v' andalso i = 0 then NONE | |
| 647 | else SOME (((v, i), x), mk ((v', 0), x))) (vs ~~ names) | |
| 648 | end; | |
| 649 | ||
| 650 | fun desymbolize_tvars thy thms = | |
| 651 | let | |
| 652 | val tvs = fold (Term.add_tvars o Thm.prop_of) thms []; | |
| 653 | val tvar_subst = mk_desymbolization (unprefix "'") (prefix "'") TVar tvs; | |
| 654 | in map (Thm.certify_instantiate (tvar_subst, [])) thms end; | |
| 655 | ||
| 656 | fun desymbolize_vars thy thm = | |
| 657 | let | |
| 658 | val vs = Term.add_vars (Thm.prop_of thm) []; | |
| 659 | val var_subst = mk_desymbolization I I Var vs; | |
| 660 | in Thm.certify_instantiate ([], var_subst) thm end; | |
| 661 | ||
| 662 | fun canonize_thms thy = desymbolize_tvars thy #> same_arity thy #> map (desymbolize_vars thy); | |
| 31156 | 663 | |
| 34874 | 664 | |
| 36112 
7fa17a225852
user interface for abstract datatypes is an attribute, not a command
 haftmann parents: 
35845diff
changeset | 665 | (* abstype certificates *) | 
| 
7fa17a225852
user interface for abstract datatypes is an attribute, not a command
 haftmann parents: 
35845diff
changeset | 666 | |
| 
7fa17a225852
user interface for abstract datatypes is an attribute, not a command
 haftmann parents: 
35845diff
changeset | 667 | fun check_abstype_cert thy proto_thm = | 
| 
7fa17a225852
user interface for abstract datatypes is an attribute, not a command
 haftmann parents: 
35845diff
changeset | 668 | let | 
| 36209 
566be5d48090
more appropriate representation of valid positions for abstractors; more accurate checking of abstype certificates
 haftmann parents: 
36202diff
changeset | 669 | val thm = (AxClass.unoverload thy o meta_rewrite thy) proto_thm; | 
| 36112 
7fa17a225852
user interface for abstract datatypes is an attribute, not a command
 haftmann parents: 
35845diff
changeset | 670 | fun bad s = bad_thm (s ^ ":\n" ^ Display.string_of_thm_global thy thm); | 
| 
7fa17a225852
user interface for abstract datatypes is an attribute, not a command
 haftmann parents: 
35845diff
changeset | 671 | val (lhs, rhs) = Logic.dest_equals (Thm.plain_prop_of thm) | 
| 
7fa17a225852
user interface for abstract datatypes is an attribute, not a command
 haftmann parents: 
35845diff
changeset | 672 | handle TERM _ => bad "Not an equation" | 
| 
7fa17a225852
user interface for abstract datatypes is an attribute, not a command
 haftmann parents: 
35845diff
changeset | 673 | | THM _ => bad "Not a proper equation"; | 
| 36209 
566be5d48090
more appropriate representation of valid positions for abstractors; more accurate checking of abstype certificates
 haftmann parents: 
36202diff
changeset | 674 | val ((abs, raw_ty), ((rep, rep_ty), param)) = (apsnd (apfst dest_Const o dest_comb) | 
| 36112 
7fa17a225852
user interface for abstract datatypes is an attribute, not a command
 haftmann parents: 
35845diff
changeset | 675 | o apfst dest_Const o dest_comb) lhs | 
| 
7fa17a225852
user interface for abstract datatypes is an attribute, not a command
 haftmann parents: 
35845diff
changeset | 676 | handle TERM _ => bad "Not an abstype certificate"; | 
| 36209 
566be5d48090
more appropriate representation of valid positions for abstractors; more accurate checking of abstype certificates
 haftmann parents: 
36202diff
changeset | 677 | val _ = pairself (fn c => if (is_some o AxClass.class_of_param thy) c | 
| 
566be5d48090
more appropriate representation of valid positions for abstractors; more accurate checking of abstype certificates
 haftmann parents: 
36202diff
changeset | 678 |       then error ("Is a class parameter: " ^ string_of_const thy c) else ()) (abs, rep);
 | 
| 
566be5d48090
more appropriate representation of valid positions for abstractors; more accurate checking of abstype certificates
 haftmann parents: 
36202diff
changeset | 679 | val _ = check_decl_ty thy (abs, raw_ty); | 
| 
566be5d48090
more appropriate representation of valid positions for abstractors; more accurate checking of abstype certificates
 haftmann parents: 
36202diff
changeset | 680 | val _ = check_decl_ty thy (rep, rep_ty); | 
| 36112 
7fa17a225852
user interface for abstract datatypes is an attribute, not a command
 haftmann parents: 
35845diff
changeset | 681 | val var = (fst o dest_Var) param | 
| 
7fa17a225852
user interface for abstract datatypes is an attribute, not a command
 haftmann parents: 
35845diff
changeset | 682 | handle TERM _ => bad "Not an abstype certificate"; | 
| 
7fa17a225852
user interface for abstract datatypes is an attribute, not a command
 haftmann parents: 
35845diff
changeset | 683 | val _ = if param = rhs then () else bad "Not an abstype certificate"; | 
| 
7fa17a225852
user interface for abstract datatypes is an attribute, not a command
 haftmann parents: 
35845diff
changeset | 684 | val ((tyco, sorts), (abs, (vs, ty'))) = ty_sorts thy (abs, Logic.unvarifyT_global raw_ty); | 
| 
7fa17a225852
user interface for abstract datatypes is an attribute, not a command
 haftmann parents: 
35845diff
changeset | 685 | val ty = domain_type ty'; | 
| 
7fa17a225852
user interface for abstract datatypes is an attribute, not a command
 haftmann parents: 
35845diff
changeset | 686 | val ty_abs = range_type ty'; | 
| 
7fa17a225852
user interface for abstract datatypes is an attribute, not a command
 haftmann parents: 
35845diff
changeset | 687 | in (tyco, (vs ~~ sorts, ((abs, ty), (rep, thm)))) end; | 
| 
7fa17a225852
user interface for abstract datatypes is an attribute, not a command
 haftmann parents: 
35845diff
changeset | 688 | |
| 
7fa17a225852
user interface for abstract datatypes is an attribute, not a command
 haftmann parents: 
35845diff
changeset | 689 | |
| 34874 | 690 | (* code equation certificates *) | 
| 691 | ||
| 34895 | 692 | fun build_head thy (c, ty) = | 
| 693 |   Thm.cterm_of thy (Logic.mk_equals (Free ("HEAD", ty), Const (c, ty)));
 | |
| 34874 | 694 | |
| 34895 | 695 | fun get_head thy cert_thm = | 
| 696 | let | |
| 697 | val [head] = (#hyps o Thm.crep_thm) cert_thm; | |
| 698 | val (_, Const (c, ty)) = (Logic.dest_equals o Thm.term_of) head; | |
| 699 | in (typscheme thy (c, ty), head) end; | |
| 700 | ||
| 35226 | 701 | fun typscheme_projection thy = | 
| 702 | typscheme thy o dest_Const o fst o dest_comb o fst o Logic.dest_equals; | |
| 703 | ||
| 704 | fun typscheme_abs thy = | |
| 705 | typscheme thy o dest_Const o fst o strip_comb o snd o dest_comb o fst o Logic.dest_equals o Thm.prop_of; | |
| 706 | ||
| 707 | fun constrain_thm thy vs sorts thm = | |
| 708 | let | |
| 709 | val mapping = map2 (fn (v, sort) => fn sort' => | |
| 710 | (v, Sorts.inter_sort (Sign.classes_of thy) (sort, sort'))) vs sorts; | |
| 711 | val inst = map2 (fn (v, sort) => fn (_, sort') => | |
| 712 | (((v, 0), sort), TFree (v, sort'))) vs mapping; | |
| 713 | val subst = (map_types o map_atyps) | |
| 714 | (fn TFree (v, _) => TFree (v, the (AList.lookup (op =) mapping v))); | |
| 715 | in | |
| 716 | thm | |
| 35845 
e5980f0ad025
renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
 wenzelm parents: 
35624diff
changeset | 717 | |> Thm.varifyT_global | 
| 35226 | 718 | |> Thm.certify_instantiate (inst, []) | 
| 719 | |> pair subst | |
| 720 | end; | |
| 721 | ||
| 722 | fun concretify_abs thy tyco abs_thm = | |
| 723 | let | |
| 724 | val (vs, ((c, _), (_, cert))) = get_abstype_spec thy tyco; | |
| 725 | val lhs = (fst o Logic.dest_equals o Thm.prop_of) abs_thm | |
| 726 | val ty = fastype_of lhs; | |
| 727 | val ty_abs = (fastype_of o snd o dest_comb) lhs; | |
| 728 | val abs = Thm.cterm_of thy (Const (c, ty --> ty_abs)); | |
| 729 | val raw_concrete_thm = Drule.transitive_thm OF [Thm.symmetric cert, Thm.combination (Thm.reflexive abs) abs_thm]; | |
| 35845 
e5980f0ad025
renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
 wenzelm parents: 
35624diff
changeset | 730 | in (c, (Thm.varifyT_global o zero_var_indexes) raw_concrete_thm) end; | 
| 35226 | 731 | |
| 732 | fun add_rhss_of_eqn thy t = | |
| 733 | let | |
| 734 | val (args, rhs) = (apfst (snd o strip_comb) o Logic.dest_equals o subst_signatures thy) t; | |
| 735 | fun add_const (Const (c, ty)) = insert (op =) (c, Sign.const_typargs thy (c, ty)) | |
| 736 | | add_const _ = I | |
| 39568 
839a0446630b
corrected long-overlooked slip: the Pure equality of a code equation is no part of the code equation itself
 haftmann parents: 
39552diff
changeset | 737 | val add_consts = fold_aterms add_const | 
| 
839a0446630b
corrected long-overlooked slip: the Pure equality of a code equation is no part of the code equation itself
 haftmann parents: 
39552diff
changeset | 738 | in add_consts rhs o fold add_consts args end; | 
| 35226 | 739 | |
| 35845 
e5980f0ad025
renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
 wenzelm parents: 
35624diff
changeset | 740 | fun dest_eqn thy = | 
| 
e5980f0ad025
renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
 wenzelm parents: 
35624diff
changeset | 741 | apfst (snd o strip_comb) o Logic.dest_equals o subst_signatures thy o Logic.unvarify_global; | 
| 35226 | 742 | |
| 743 | abstype cert = Equations of thm * bool list | |
| 744 | | Projection of term * string | |
| 745 | | Abstract of thm * string | |
| 746 | with | |
| 34891 
99b9a6290446
code certificates as integral part of code generation
 haftmann parents: 
34874diff
changeset | 747 | |
| 
99b9a6290446
code certificates as integral part of code generation
 haftmann parents: 
34874diff
changeset | 748 | fun empty_cert thy c = | 
| 
99b9a6290446
code certificates as integral part of code generation
 haftmann parents: 
34874diff
changeset | 749 | let | 
| 
99b9a6290446
code certificates as integral part of code generation
 haftmann parents: 
34874diff
changeset | 750 | val raw_ty = const_typ thy c; | 
| 
99b9a6290446
code certificates as integral part of code generation
 haftmann parents: 
34874diff
changeset | 751 | val tvars = Term.add_tvar_namesT raw_ty []; | 
| 
99b9a6290446
code certificates as integral part of code generation
 haftmann parents: 
34874diff
changeset | 752 | val tvars' = case AxClass.class_of_param thy c | 
| 
99b9a6290446
code certificates as integral part of code generation
 haftmann parents: 
34874diff
changeset | 753 | of SOME class => [TFree (Name.aT, [class])] | 
| 
99b9a6290446
code certificates as integral part of code generation
 haftmann parents: 
34874diff
changeset | 754 | | NONE => Name.invent_list [] Name.aT (length tvars) | 
| 
99b9a6290446
code certificates as integral part of code generation
 haftmann parents: 
34874diff
changeset | 755 | |> map (fn v => TFree (v, [])); | 
| 
99b9a6290446
code certificates as integral part of code generation
 haftmann parents: 
34874diff
changeset | 756 | val ty = typ_subst_TVars (tvars ~~ tvars') raw_ty; | 
| 34895 | 757 | val chead = build_head thy (c, ty); | 
| 35226 | 758 | in Equations (Thm.weaken chead Drule.dummy_thm, []) end; | 
| 34891 
99b9a6290446
code certificates as integral part of code generation
 haftmann parents: 
34874diff
changeset | 759 | |
| 
99b9a6290446
code certificates as integral part of code generation
 haftmann parents: 
34874diff
changeset | 760 | fun cert_of_eqns thy c [] = empty_cert thy c | 
| 34895 | 761 | | cert_of_eqns thy c raw_eqns = | 
| 34874 | 762 | let | 
| 34895 | 763 | val eqns = burrow_fst (canonize_thms thy) raw_eqns; | 
| 764 | val _ = map (assert_eqn thy) eqns; | |
| 34891 
99b9a6290446
code certificates as integral part of code generation
 haftmann parents: 
34874diff
changeset | 765 | val (thms, propers) = split_list eqns; | 
| 34895 | 766 | val _ = map (fn thm => if c = const_eqn thy thm then () | 
| 767 |           else error ("Wrong head of code equation,\nexpected constant "
 | |
| 768 | ^ string_of_const thy c ^ "\n" ^ Display.string_of_thm_global thy thm)) thms; | |
| 34891 
99b9a6290446
code certificates as integral part of code generation
 haftmann parents: 
34874diff
changeset | 769 | fun tvars_of T = rev (Term.add_tvarsT T []); | 
| 
99b9a6290446
code certificates as integral part of code generation
 haftmann parents: 
34874diff
changeset | 770 | val vss = map (tvars_of o snd o head_eqn) thms; | 
| 
99b9a6290446
code certificates as integral part of code generation
 haftmann parents: 
34874diff
changeset | 771 | fun inter_sorts vs = | 
| 
99b9a6290446
code certificates as integral part of code generation
 haftmann parents: 
34874diff
changeset | 772 | fold (curry (Sorts.inter_sort (Sign.classes_of thy)) o snd) vs []; | 
| 
99b9a6290446
code certificates as integral part of code generation
 haftmann parents: 
34874diff
changeset | 773 | val sorts = map_transpose inter_sorts vss; | 
| 
99b9a6290446
code certificates as integral part of code generation
 haftmann parents: 
34874diff
changeset | 774 | val vts = Name.names Name.context Name.aT sorts; | 
| 
99b9a6290446
code certificates as integral part of code generation
 haftmann parents: 
34874diff
changeset | 775 | val thms as thm :: _ = | 
| 
99b9a6290446
code certificates as integral part of code generation
 haftmann parents: 
34874diff
changeset | 776 | map2 (fn vs => Thm.certify_instantiate (vs ~~ map TFree vts, [])) vss thms; | 
| 34895 | 777 | val head_thm = Thm.symmetric (Thm.assume (build_head thy (head_eqn (hd thms)))); | 
| 34874 | 778 | fun head_conv ct = if can Thm.dest_comb ct | 
| 779 | then Conv.fun_conv head_conv ct | |
| 780 | else Conv.rewr_conv head_thm ct; | |
| 781 | val rewrite_head = Conv.fconv_rule (Conv.arg1_conv head_conv); | |
| 34891 
99b9a6290446
code certificates as integral part of code generation
 haftmann parents: 
34874diff
changeset | 782 | val cert_thm = Conjunction.intr_balanced (map rewrite_head thms); | 
| 35226 | 783 | in Equations (cert_thm, propers) end; | 
| 34891 
99b9a6290446
code certificates as integral part of code generation
 haftmann parents: 
34874diff
changeset | 784 | |
| 35226 | 785 | fun cert_of_proj thy c tyco = | 
| 786 | let | |
| 787 | val (vs, ((abs, ty), (rep, cert))) = get_abstype_spec thy tyco; | |
| 788 | val _ = if c = rep then () else | |
| 789 |       error ("Wrong head of projection,\nexpected constant " ^ string_of_const thy rep);
 | |
| 790 | in Projection (mk_proj tyco vs ty abs rep, tyco) end; | |
| 791 | ||
| 792 | fun cert_of_abs thy tyco c raw_abs_thm = | |
| 34874 | 793 | let | 
| 35226 | 794 | val abs_thm = singleton (canonize_thms thy) raw_abs_thm; | 
| 795 | val _ = assert_abs_eqn thy (SOME tyco) abs_thm; | |
| 796 | val _ = if c = const_abs_eqn thy abs_thm then () | |
| 797 |       else error ("Wrong head of abstract code equation,\nexpected constant "
 | |
| 798 | ^ string_of_const thy c ^ "\n" ^ Display.string_of_thm_global thy abs_thm); | |
| 36615 
88756a5a92fc
renamed Thm.freezeT to Thm.legacy_freezeT -- it is based on Type.legacy_freeze;
 wenzelm parents: 
36610diff
changeset | 799 | in Abstract (Thm.legacy_freezeT abs_thm, tyco) end; | 
| 34874 | 800 | |
| 35226 | 801 | fun constrain_cert thy sorts (Equations (cert_thm, propers)) = | 
| 802 | let | |
| 803 | val ((vs, _), head) = get_head thy cert_thm; | |
| 804 | val (subst, cert_thm') = cert_thm | |
| 805 | |> Thm.implies_intr head | |
| 806 | |> constrain_thm thy vs sorts; | |
| 807 | val head' = Thm.term_of head | |
| 808 | |> subst | |
| 809 | |> Thm.cterm_of thy; | |
| 810 | val cert_thm'' = cert_thm' | |
| 811 | |> Thm.elim_implies (Thm.assume head'); | |
| 812 | in Equations (cert_thm'', propers) end | |
| 813 | | constrain_cert thy _ (cert as Projection _) = | |
| 814 | cert | |
| 815 | | constrain_cert thy sorts (Abstract (abs_thm, tyco)) = | |
| 816 | Abstract (snd (constrain_thm thy (fst (typscheme_abs thy abs_thm)) sorts abs_thm), tyco); | |
| 817 | ||
| 818 | fun typscheme_of_cert thy (Equations (cert_thm, _)) = | |
| 819 | fst (get_head thy cert_thm) | |
| 820 | | typscheme_of_cert thy (Projection (proj, _)) = | |
| 821 | typscheme_projection thy proj | |
| 822 | | typscheme_of_cert thy (Abstract (abs_thm, _)) = | |
| 823 | typscheme_abs thy abs_thm; | |
| 34874 | 824 | |
| 35226 | 825 | fun typargs_deps_of_cert thy (Equations (cert_thm, propers)) = | 
| 826 | let | |
| 827 | val vs = (fst o fst) (get_head thy cert_thm); | |
| 828 | val equations = if null propers then [] else | |
| 829 | Thm.prop_of cert_thm | |
| 830 | |> Logic.dest_conjunction_balanced (length propers); | |
| 831 | in (vs, fold (add_rhss_of_eqn thy) equations []) end | |
| 832 | | typargs_deps_of_cert thy (Projection (t, tyco)) = | |
| 833 | (fst (typscheme_projection thy t), add_rhss_of_eqn thy t []) | |
| 834 | | typargs_deps_of_cert thy (Abstract (abs_thm, tyco)) = | |
| 835 | let | |
| 836 | val vs = fst (typscheme_abs thy abs_thm); | |
| 837 | val (_, concrete_thm) = concretify_abs thy tyco abs_thm; | |
| 36202 
43ea1f28fc7c
explicit check sorts in abstract certificates; abstractor is part of dependencies
 haftmann parents: 
36122diff
changeset | 838 | in (vs, add_rhss_of_eqn thy (map_types Logic.unvarifyT_global (Thm.prop_of concrete_thm)) []) end; | 
| 34895 | 839 | |
| 35226 | 840 | fun equations_of_cert thy (cert as Equations (cert_thm, propers)) = | 
| 841 | let | |
| 842 | val tyscm = typscheme_of_cert thy cert; | |
| 843 | val thms = if null propers then [] else | |
| 844 | cert_thm | |
| 35624 | 845 | |> Local_Defs.expand [snd (get_head thy cert_thm)] | 
| 35845 
e5980f0ad025
renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
 wenzelm parents: 
35624diff
changeset | 846 | |> Thm.varifyT_global | 
| 35226 | 847 | |> Conjunction.elim_balanced (length propers); | 
| 36209 
566be5d48090
more appropriate representation of valid positions for abstractors; more accurate checking of abstype certificates
 haftmann parents: 
36202diff
changeset | 848 | fun abstractions (args, rhs) = (map (rpair NONE) args, (rhs, NONE)); | 
| 
566be5d48090
more appropriate representation of valid positions for abstractors; more accurate checking of abstype certificates
 haftmann parents: 
36202diff
changeset | 849 | in (tyscm, map (abstractions o dest_eqn thy o Thm.prop_of) thms ~~ (map SOME thms ~~ propers)) end | 
| 35226 | 850 | | equations_of_cert thy (Projection (t, tyco)) = | 
| 851 | let | |
| 852 | val (_, ((abs, _), _)) = get_abstype_spec thy tyco; | |
| 853 | val tyscm = typscheme_projection thy t; | |
| 35845 
e5980f0ad025
renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
 wenzelm parents: 
35624diff
changeset | 854 | val t' = map_types Logic.varifyT_global t; | 
| 36209 
566be5d48090
more appropriate representation of valid positions for abstractors; more accurate checking of abstype certificates
 haftmann parents: 
36202diff
changeset | 855 | fun abstractions (args, rhs) = (map (rpair (SOME abs)) args, (rhs, NONE)); | 
| 
566be5d48090
more appropriate representation of valid positions for abstractors; more accurate checking of abstype certificates
 haftmann parents: 
36202diff
changeset | 856 | in (tyscm, [((abstractions o dest_eqn thy) t', (NONE, true))]) end | 
| 35226 | 857 | | equations_of_cert thy (Abstract (abs_thm, tyco)) = | 
| 858 | let | |
| 859 | val tyscm = typscheme_abs thy abs_thm; | |
| 860 | val (abs, concrete_thm) = concretify_abs thy tyco abs_thm; | |
| 36209 
566be5d48090
more appropriate representation of valid positions for abstractors; more accurate checking of abstype certificates
 haftmann parents: 
36202diff
changeset | 861 | fun abstractions (args, rhs) = (map (rpair NONE) args, (rhs, (SOME abs))); | 
| 35845 
e5980f0ad025
renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
 wenzelm parents: 
35624diff
changeset | 862 | in | 
| 36209 
566be5d48090
more appropriate representation of valid positions for abstractors; more accurate checking of abstype certificates
 haftmann parents: 
36202diff
changeset | 863 | (tyscm, [((abstractions o dest_eqn thy o Thm.prop_of) concrete_thm, | 
| 
566be5d48090
more appropriate representation of valid positions for abstractors; more accurate checking of abstype certificates
 haftmann parents: 
36202diff
changeset | 864 | (SOME (Thm.varifyT_global abs_thm), true))]) | 
| 35845 
e5980f0ad025
renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
 wenzelm parents: 
35624diff
changeset | 865 | end; | 
| 34895 | 866 | |
| 35226 | 867 | fun pretty_cert thy (cert as Equations _) = | 
| 868 | (map_filter (Option.map (Display.pretty_thm_global thy o AxClass.overload thy) o fst o snd) | |
| 869 | o snd o equations_of_cert thy) cert | |
| 870 | | pretty_cert thy (Projection (t, _)) = | |
| 35845 
e5980f0ad025
renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
 wenzelm parents: 
35624diff
changeset | 871 | [Syntax.pretty_term_global thy (map_types Logic.varifyT_global t)] | 
| 35226 | 872 | | pretty_cert thy (Abstract (abs_thm, tyco)) = | 
| 35845 
e5980f0ad025
renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
 wenzelm parents: 
35624diff
changeset | 873 | [(Display.pretty_thm_global thy o AxClass.overload thy o Thm.varifyT_global) abs_thm]; | 
| 35226 | 874 | |
| 875 | fun bare_thms_of_cert thy (cert as Equations _) = | |
| 876 | (map_filter (fn (_, (some_thm, proper)) => if proper then some_thm else NONE) | |
| 877 | o snd o equations_of_cert thy) cert | |
| 35376 
212b1dc5212d
use abstract code cerficates for bare code theorems
 haftmann parents: 
35373diff
changeset | 878 | | bare_thms_of_cert thy (Projection _) = [] | 
| 
212b1dc5212d
use abstract code cerficates for bare code theorems
 haftmann parents: 
35373diff
changeset | 879 | | bare_thms_of_cert thy (Abstract (abs_thm, tyco)) = | 
| 35845 
e5980f0ad025
renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
 wenzelm parents: 
35624diff
changeset | 880 | [Thm.varifyT_global (snd (concretify_abs thy tyco abs_thm))]; | 
| 34895 | 881 | |
| 882 | end; | |
| 34891 
99b9a6290446
code certificates as integral part of code generation
 haftmann parents: 
34874diff
changeset | 883 | |
| 34874 | 884 | |
| 35226 | 885 | (* code certificate access *) | 
| 886 | ||
| 887 | fun retrieve_raw thy c = | |
| 888 | Symtab.lookup ((the_functions o the_exec) thy) c | |
| 889 | |> Option.map (snd o fst) | |
| 37460 
910b2422571d
drop subsumed default equations (requires a little bit unfortunate laziness)
 haftmann parents: 
37448diff
changeset | 890 | |> the_default empty_fun_spec | 
| 34874 | 891 | |
| 35226 | 892 | fun get_cert thy f c = case retrieve_raw thy c | 
| 37460 
910b2422571d
drop subsumed default equations (requires a little bit unfortunate laziness)
 haftmann parents: 
37448diff
changeset | 893 | of Default (_, eqns_lazy) => Lazy.force eqns_lazy | 
| 35226 | 894 | |> (map o apfst) (Thm.transfer thy) | 
| 895 | |> f | |
| 896 | |> (map o apfst) (AxClass.unoverload thy) | |
| 897 | |> cert_of_eqns thy c | |
| 898 | | Eqns eqns => eqns | |
| 899 | |> (map o apfst) (Thm.transfer thy) | |
| 900 | |> f | |
| 901 | |> (map o apfst) (AxClass.unoverload thy) | |
| 902 | |> cert_of_eqns thy c | |
| 903 | | Proj (_, tyco) => | |
| 904 | cert_of_proj thy c tyco | |
| 905 | | Abstr (abs_thm, tyco) => abs_thm | |
| 906 | |> Thm.transfer thy | |
| 907 | |> AxClass.unoverload thy | |
| 908 | |> cert_of_abs thy tyco c; | |
| 31962 | 909 | |
| 910 | ||
| 911 | (* cases *) | |
| 31156 | 912 | |
| 913 | fun case_certificate thm = | |
| 914 | let | |
| 915 | 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 | 916 | o apsnd Logic.dest_conjunctions o Logic.dest_implies o Thm.plain_prop_of) thm; | 
| 31156 | 917 | val _ = case head of Free _ => true | 
| 918 | | Var _ => true | |
| 919 |       | _ => raise TERM ("case_cert", []);
 | |
| 920 | val ([(case_var, _)], case_expr) = Term.strip_abs_eta 1 raw_case_expr; | |
| 921 | val (Const (case_const, _), raw_params) = strip_comb case_expr; | |
| 922 | val n = find_index (fn Free (v, _) => v = case_var | _ => false) raw_params; | |
| 923 |     val _ = if n = ~1 then raise TERM ("case_cert", []) else ();
 | |
| 924 | val params = map (fst o dest_Var) (nth_drop n raw_params); | |
| 925 | fun dest_case t = | |
| 926 | let | |
| 927 | val (head' $ t_co, rhs) = Logic.dest_equals t; | |
| 928 |         val _ = if head' = head then () else raise TERM ("case_cert", []);
 | |
| 929 | val (Const (co, _), args) = strip_comb t_co; | |
| 930 | val (Var (param, _), args') = strip_comb rhs; | |
| 931 |         val _ = if args' = args then () else raise TERM ("case_cert", []);
 | |
| 932 | in (param, co) end; | |
| 933 | fun analyze_cases cases = | |
| 934 | let | |
| 935 | val co_list = fold (AList.update (op =) o dest_case) cases []; | |
| 936 | in map (the o AList.lookup (op =) co_list) params end; | |
| 937 | fun analyze_let t = | |
| 938 | let | |
| 939 | val (head' $ arg, Var (param', _) $ arg') = Logic.dest_equals t; | |
| 940 |         val _ = if head' = head then () else raise TERM ("case_cert", []);
 | |
| 941 |         val _ = if arg' = arg then () else raise TERM ("case_cert", []);
 | |
| 942 |         val _ = if [param'] = params then () else raise TERM ("case_cert", []);
 | |
| 943 | in [] end; | |
| 944 | fun analyze (cases as [let_case]) = | |
| 945 | (analyze_cases cases handle Bind => analyze_let let_case) | |
| 946 | | analyze cases = analyze_cases cases; | |
| 947 | in (case_const, (n, analyze cases)) end; | |
| 948 | ||
| 949 | fun case_cert thm = case_certificate thm | |
| 950 | handle Bind => error "bad case certificate" | |
| 951 | | TERM _ => error "bad case certificate"; | |
| 952 | ||
| 37438 | 953 | fun get_case_scheme thy = Option.map fst o Symtab.lookup ((fst o the_cases o the_exec) thy); | 
| 954 | fun get_case_cong thy = Option.map snd o Symtab.lookup ((fst o the_cases o the_exec) thy); | |
| 24219 | 955 | |
| 31962 | 956 | val undefineds = Symtab.keys o snd o the_cases o the_exec; | 
| 24219 | 957 | |
| 958 | ||
| 31962 | 959 | (* diagnostic *) | 
| 24219 | 960 | |
| 961 | fun print_codesetup thy = | |
| 962 | let | |
| 36610 
bafd82950e24
renamed ProofContext.init to ProofContext.init_global to emphasize that this is not the real thing;
 wenzelm parents: 
36209diff
changeset | 963 | val ctxt = ProofContext.init_global thy; | 
| 24844 
98c006a30218
certificates for code generator case expressions
 haftmann parents: 
24837diff
changeset | 964 | val exec = the_exec thy; | 
| 35226 | 965 | fun pretty_equations const thms = | 
| 24219 | 966 | (Pretty.block o Pretty.fbreaks) ( | 
| 35226 | 967 | Pretty.str (string_of_const thy const) :: map (Display.pretty_thm ctxt) thms | 
| 24219 | 968 | ); | 
| 37460 
910b2422571d
drop subsumed default equations (requires a little bit unfortunate laziness)
 haftmann parents: 
37448diff
changeset | 969 | fun pretty_function (const, Default (_, eqns_lazy)) = pretty_equations const (map fst (Lazy.force eqns_lazy)) | 
| 35226 | 970 | | pretty_function (const, Eqns eqns) = pretty_equations const (map fst eqns) | 
| 971 | | pretty_function (const, Proj (proj, _)) = Pretty.block | |
| 972 | [Pretty.str (string_of_const thy const), Pretty.fbrk, Syntax.pretty_term ctxt proj] | |
| 973 | | pretty_function (const, Abstr (thm, _)) = pretty_equations const [thm]; | |
| 974 | fun pretty_typ (tyco, vs) = Pretty.str | |
| 975 | (string_of_typ thy (Type (tyco, map TFree vs))); | |
| 976 | fun pretty_typspec (typ, (cos, abstract)) = if null cos | |
| 977 | then pretty_typ typ | |
| 978 | else (Pretty.block o Pretty.breaks) ( | |
| 979 | pretty_typ typ | |
| 980 | :: Pretty.str "=" | |
| 981 | :: (if abstract then [Pretty.str "(abstract)"] else []) | |
| 982 | @ separate (Pretty.str "|") (map (fn (c, []) => Pretty.str (string_of_const thy c) | |
| 983 | | (c, tys) => | |
| 984 | (Pretty.block o Pretty.breaks) | |
| 985 | (Pretty.str (string_of_const thy c) | |
| 986 | :: Pretty.str "of" | |
| 987 | :: map (Pretty.quote o Syntax.pretty_typ_global thy) tys)) cos) | |
| 988 | ); | |
| 37438 | 989 | fun pretty_case (const, ((_, (_, [])), _)) = Pretty.str (string_of_const thy const) | 
| 990 | | pretty_case (const, ((_, (_, cos)), _)) = (Pretty.block o Pretty.breaks) [ | |
| 34901 | 991 | Pretty.str (string_of_const thy const), Pretty.str "with", | 
| 992 | (Pretty.block o Pretty.commas o map (Pretty.str o string_of_const thy)) cos]; | |
| 35226 | 993 | val functions = the_functions exec | 
| 24423 
ae9cd0e92423
overloaded definitions accompanied by explicit constants
 haftmann parents: 
24283diff
changeset | 994 | |> Symtab.dest | 
| 28695 | 995 | |> (map o apsnd) (snd o fst) | 
| 24219 | 996 | |> sort (string_ord o pairself fst); | 
| 35299 
4f4d5bf4ea08
proper distinction of code datatypes and abstypes
 haftmann parents: 
35226diff
changeset | 997 | val datatypes = the_types exec | 
| 24219 | 998 | |> Symtab.dest | 
| 35226 | 999 | |> map (fn (tyco, (_, (vs, spec)) :: _) => | 
| 1000 | ((tyco, vs), constructors_of spec)) | |
| 1001 | |> sort (string_ord o pairself (fst o fst)); | |
| 34901 | 1002 | val cases = Symtab.dest ((fst o the_cases o the_exec) thy); | 
| 1003 | val undefineds = Symtab.keys ((snd o the_cases o the_exec) thy); | |
| 24219 | 1004 | in | 
| 1005 | (Pretty.writeln o Pretty.chunks) [ | |
| 1006 | Pretty.block ( | |
| 34901 | 1007 | Pretty.str "code equations:" :: Pretty.fbrk | 
| 35226 | 1008 | :: (Pretty.fbreaks o map pretty_function) functions | 
| 24219 | 1009 | ), | 
| 25968 | 1010 | Pretty.block ( | 
| 34901 | 1011 | Pretty.str "datatypes:" :: Pretty.fbrk | 
| 35226 | 1012 | :: (Pretty.fbreaks o map pretty_typspec) datatypes | 
| 34901 | 1013 | ), | 
| 1014 | Pretty.block ( | |
| 1015 | Pretty.str "cases:" :: Pretty.fbrk | |
| 1016 | :: (Pretty.fbreaks o map pretty_case) cases | |
| 1017 | ), | |
| 1018 | Pretty.block ( | |
| 1019 | Pretty.str "undefined:" :: Pretty.fbrk | |
| 1020 | :: (Pretty.commas o map (Pretty.str o string_of_const thy)) undefineds | |
| 24219 | 1021 | ) | 
| 1022 | ] | |
| 1023 | end; | |
| 1024 | ||
| 1025 | ||
| 31962 | 1026 | (** declaring executable ingredients **) | 
| 1027 | ||
| 33940 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 1028 | (* constant signatures *) | 
| 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 1029 | |
| 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 1030 | fun add_type tyco thy = | 
| 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 1031 | case Symtab.lookup ((snd o #types o Type.rep_tsig o Sign.tsig_of) thy) tyco | 
| 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 1032 | of SOME (Type.Abbreviation (vs, _, _)) => | 
| 34173 
458ced35abb8
reduced code generator cache to the baremost minimum
 haftmann parents: 
33977diff
changeset | 1033 | (map_exec_purge o map_signatures o apfst) | 
| 33940 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 1034 | (Symtab.update (tyco, length vs)) thy | 
| 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 1035 |     | _ => error ("No such type abbreviation: " ^ quote tyco);
 | 
| 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 1036 | |
| 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 1037 | fun add_type_cmd s thy = add_type (Sign.intern_type thy s) thy; | 
| 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 1038 | |
| 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 1039 | fun gen_add_signature prep_const prep_signature (raw_c, raw_ty) thy = | 
| 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 1040 | let | 
| 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 1041 | val c = prep_const thy raw_c; | 
| 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 1042 | val ty = prep_signature thy raw_ty; | 
| 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 1043 | val ty' = expand_signature thy ty; | 
| 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 1044 | val ty'' = Sign.the_const_type thy c; | 
| 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 1045 | val _ = if typ_equiv (ty', ty'') then () else | 
| 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 1046 |       error ("Illegal constant signature: " ^ Syntax.string_of_typ_global thy ty);
 | 
| 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 1047 | in | 
| 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 1048 | thy | 
| 34173 
458ced35abb8
reduced code generator cache to the baremost minimum
 haftmann parents: 
33977diff
changeset | 1049 | |> (map_exec_purge o map_signatures o apsnd) (Symtab.update (c, ty)) | 
| 33940 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 1050 | end; | 
| 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 1051 | |
| 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 1052 | val add_signature = gen_add_signature (K I) cert_signature; | 
| 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 1053 | val add_signature_cmd = gen_add_signature read_const read_signature; | 
| 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 1054 | |
| 
317933ce3712
crude support for type aliasses and corresponding constant signatures
 haftmann parents: 
33756diff
changeset | 1055 | |
| 31962 | 1056 | (* code equations *) | 
| 1057 | ||
| 35226 | 1058 | fun gen_add_eqn default (raw_thm, proper) thy = | 
| 33075 | 1059 | let | 
| 35226 | 1060 | val thm = Thm.close_derivation raw_thm; | 
| 1061 | val c = const_eqn thy thm; | |
| 37460 
910b2422571d
drop subsumed default equations (requires a little bit unfortunate laziness)
 haftmann parents: 
37448diff
changeset | 1062 | fun update_subsume thy (thm, proper) eqns = | 
| 
910b2422571d
drop subsumed default equations (requires a little bit unfortunate laziness)
 haftmann parents: 
37448diff
changeset | 1063 | let | 
| 39811 | 1064 | val args_of = snd o chop_while is_Var o rev o snd o strip_comb | 
| 39791 
a91430778479
redundancy check: drop trailing Var arguments (avoids eta problems with equations)
 haftmann parents: 
39568diff
changeset | 1065 | o map_types Type.strip_sorts o fst o Logic.dest_equals o Thm.plain_prop_of; | 
| 37460 
910b2422571d
drop subsumed default equations (requires a little bit unfortunate laziness)
 haftmann parents: 
37448diff
changeset | 1066 | val args = args_of thm; | 
| 
910b2422571d
drop subsumed default equations (requires a little bit unfortunate laziness)
 haftmann parents: 
37448diff
changeset | 1067 | val incr_idx = Logic.incr_indexes ([], Thm.maxidx_of thm + 1); | 
| 39794 
51451d73c3d4
corrected subsumption check: arguments have been reversed; addition of default equations to non-default equations is identity
 haftmann parents: 
39791diff
changeset | 1068 | fun matches_args args' = | 
| 
51451d73c3d4
corrected subsumption check: arguments have been reversed; addition of default equations to non-default equations is identity
 haftmann parents: 
39791diff
changeset | 1069 | let | 
| 
51451d73c3d4
corrected subsumption check: arguments have been reversed; addition of default equations to non-default equations is identity
 haftmann parents: 
39791diff
changeset | 1070 | val k = length args' - length args | 
| 
51451d73c3d4
corrected subsumption check: arguments have been reversed; addition of default equations to non-default equations is identity
 haftmann parents: 
39791diff
changeset | 1071 | in if k >= 0 | 
| 
51451d73c3d4
corrected subsumption check: arguments have been reversed; addition of default equations to non-default equations is identity
 haftmann parents: 
39791diff
changeset | 1072 | then Pattern.matchess thy (args, (map incr_idx o drop k) args') | 
| 
51451d73c3d4
corrected subsumption check: arguments have been reversed; addition of default equations to non-default equations is identity
 haftmann parents: 
39791diff
changeset | 1073 | else false | 
| 
51451d73c3d4
corrected subsumption check: arguments have been reversed; addition of default equations to non-default equations is identity
 haftmann parents: 
39791diff
changeset | 1074 | end; | 
| 37460 
910b2422571d
drop subsumed default equations (requires a little bit unfortunate laziness)
 haftmann parents: 
37448diff
changeset | 1075 | fun drop (thm', proper') = if (proper orelse not proper') | 
| 
910b2422571d
drop subsumed default equations (requires a little bit unfortunate laziness)
 haftmann parents: 
37448diff
changeset | 1076 | andalso matches_args (args_of thm') then | 
| 
910b2422571d
drop subsumed default equations (requires a little bit unfortunate laziness)
 haftmann parents: 
37448diff
changeset | 1077 |             (warning ("Code generator: dropping subsumed code equation\n" ^
 | 
| 
910b2422571d
drop subsumed default equations (requires a little bit unfortunate laziness)
 haftmann parents: 
37448diff
changeset | 1078 | Display.string_of_thm_global thy thm'); true) | 
| 
910b2422571d
drop subsumed default equations (requires a little bit unfortunate laziness)
 haftmann parents: 
37448diff
changeset | 1079 | else false; | 
| 
910b2422571d
drop subsumed default equations (requires a little bit unfortunate laziness)
 haftmann parents: 
37448diff
changeset | 1080 | in (thm, proper) :: filter_out drop eqns end; | 
| 
910b2422571d
drop subsumed default equations (requires a little bit unfortunate laziness)
 haftmann parents: 
37448diff
changeset | 1081 | fun natural_order thy_ref eqns = | 
| 
910b2422571d
drop subsumed default equations (requires a little bit unfortunate laziness)
 haftmann parents: 
37448diff
changeset | 1082 | (eqns, Lazy.lazy (fn () => fold (update_subsume (Theory.deref thy_ref)) eqns [])) | 
| 
910b2422571d
drop subsumed default equations (requires a little bit unfortunate laziness)
 haftmann parents: 
37448diff
changeset | 1083 | fun add_eqn' true (Default (eqns, _)) = | 
| 
910b2422571d
drop subsumed default equations (requires a little bit unfortunate laziness)
 haftmann parents: 
37448diff
changeset | 1084 | Default (natural_order (Theory.check_thy thy) ((thm, proper) :: eqns)) | 
| 
910b2422571d
drop subsumed default equations (requires a little bit unfortunate laziness)
 haftmann parents: 
37448diff
changeset | 1085 | (*this restores the natural order and drops syntactic redundancies*) | 
| 39794 
51451d73c3d4
corrected subsumption check: arguments have been reversed; addition of default equations to non-default equations is identity
 haftmann parents: 
39791diff
changeset | 1086 | | add_eqn' true fun_spec = fun_spec | 
| 
51451d73c3d4
corrected subsumption check: arguments have been reversed; addition of default equations to non-default equations is identity
 haftmann parents: 
39791diff
changeset | 1087 | | add_eqn' false (Eqns eqns) = Eqns (update_subsume thy (thm, proper) eqns) | 
| 35226 | 1088 | | add_eqn' false _ = Eqns [(thm, proper)]; | 
| 1089 | in change_fun_spec false c (add_eqn' default) thy end; | |
| 31962 | 1090 | |
| 1091 | fun add_eqn thm thy = | |
| 1092 | gen_add_eqn false (mk_eqn thy (thm, true)) thy; | |
| 1093 | ||
| 1094 | fun add_warning_eqn thm thy = | |
| 1095 | case mk_eqn_warning thy thm | |
| 1096 | of SOME eqn => gen_add_eqn false eqn thy | |
| 1097 | | NONE => thy; | |
| 1098 | ||
| 37425 | 1099 | fun add_nbe_eqn thm thy = | 
| 1100 | gen_add_eqn false (mk_eqn thy (thm, false)) thy; | |
| 1101 | ||
| 31962 | 1102 | fun add_default_eqn thm thy = | 
| 1103 | case mk_eqn_liberal thy thm | |
| 1104 | of SOME eqn => gen_add_eqn true eqn thy | |
| 1105 | | NONE => thy; | |
| 1106 | ||
| 1107 | val add_default_eqn_attribute = Thm.declaration_attribute | |
| 1108 | (fn thm => Context.mapping (add_default_eqn thm) I); | |
| 1109 | val add_default_eqn_attrib = Attrib.internal (K add_default_eqn_attribute); | |
| 1110 | ||
| 37425 | 1111 | fun add_nbe_default_eqn thm thy = | 
| 1112 | gen_add_eqn true (mk_eqn thy (thm, false)) thy; | |
| 1113 | ||
| 1114 | val add_nbe_default_eqn_attribute = Thm.declaration_attribute | |
| 1115 | (fn thm => Context.mapping (add_nbe_default_eqn thm) I); | |
| 1116 | val add_nbe_default_eqn_attrib = Attrib.internal (K add_nbe_default_eqn_attribute); | |
| 1117 | ||
| 35226 | 1118 | fun add_abs_eqn raw_thm thy = | 
| 1119 | let | |
| 1120 | val (abs_thm, tyco) = (apfst Thm.close_derivation o mk_abs_eqn thy) raw_thm; | |
| 1121 | val c = const_abs_eqn thy abs_thm; | |
| 1122 | in change_fun_spec false c (K (Abstr (abs_thm, tyco))) thy end; | |
| 1123 | ||
| 31962 | 1124 | fun del_eqn thm thy = case mk_eqn_liberal thy thm | 
| 35226 | 1125 | of SOME (thm, _) => let | 
| 1126 | fun del_eqn' (Default eqns) = empty_fun_spec | |
| 1127 | | del_eqn' (Eqns eqns) = | |
| 1128 | Eqns (filter_out (fn (thm', _) => Thm.eq_thm_prop (thm, thm')) eqns) | |
| 1129 | | del_eqn' spec = spec | |
| 1130 | in change_fun_spec true (const_eqn thy thm) del_eqn' thy end | |
| 31962 | 1131 | | NONE => thy; | 
| 1132 | ||
| 35226 | 1133 | fun del_eqns c = change_fun_spec true c (K empty_fun_spec); | 
| 34244 | 1134 | |
| 1135 | ||
| 1136 | (* cases *) | |
| 1137 | ||
| 37438 | 1138 | fun case_cong thy case_const (num_args, (pos, constrs)) = | 
| 1139 | let | |
| 1140 | val ([x, y], ctxt) = Name.variants ["A", "A'"] Name.context; | |
| 1141 | val (zs, _) = Name.variants (replicate (num_args - 1) "") ctxt; | |
| 1142 | val (ws, vs) = chop pos zs; | |
| 37448 
3bd4b3809bee
explicit type variable arguments for constructors
 haftmann parents: 
37438diff
changeset | 1143 | val T = Logic.unvarifyT_global (Sign.the_const_type thy case_const); | 
| 37438 | 1144 | val Ts = (fst o strip_type) T; | 
| 1145 | val T_cong = nth Ts pos; | |
| 1146 | fun mk_prem z = Free (z, T_cong); | |
| 1147 | fun mk_concl z = list_comb (Const (case_const, T), map2 (curry Free) (ws @ z :: vs) Ts); | |
| 1148 | val (prem, concl) = pairself Logic.mk_equals (pairself mk_prem (x, y), pairself mk_concl (x, y)); | |
| 37521 | 1149 |     fun tac { context, prems } = Simplifier.rewrite_goals_tac prems
 | 
| 37438 | 1150 | THEN ALLGOALS (ProofContext.fact_tac [Drule.reflexive_thm]); | 
| 1151 | in Skip_Proof.prove_global thy (x :: y :: zs) [prem] concl tac end; | |
| 1152 | ||
| 34244 | 1153 | fun add_case thm thy = | 
| 1154 | let | |
| 37438 | 1155 | val (case_const, (k, case_pats)) = case_cert thm; | 
| 34244 | 1156 | val _ = case filter_out (is_constr thy) case_pats | 
| 1157 | of [] => () | |
| 1158 |       | cs => error ("Non-constructor(s) in case certificate: " ^ commas (map quote cs));
 | |
| 37438 | 1159 | val entry = (1 + Int.max (1, length case_pats), (k, case_pats)); | 
| 1160 | in | |
| 1161 | thy | |
| 1162 | |> Theory.checkpoint | |
| 1163 | |> `(fn thy => case_cong thy case_const entry) | |
| 1164 | |-> (fn cong => (map_exec_purge o map_cases o apfst) (Symtab.update (case_const, (entry, cong)))) | |
| 1165 | end; | |
| 34244 | 1166 | |
| 1167 | fun add_undefined c thy = | |
| 1168 | (map_exec_purge o map_cases o apsnd) (Symtab.update (c, ())) thy; | |
| 1169 | ||
| 1170 | ||
| 35299 
4f4d5bf4ea08
proper distinction of code datatypes and abstypes
 haftmann parents: 
35226diff
changeset | 1171 | (* types *) | 
| 34244 | 1172 | |
| 35299 
4f4d5bf4ea08
proper distinction of code datatypes and abstypes
 haftmann parents: 
35226diff
changeset | 1173 | fun register_type (tyco, vs_spec) thy = | 
| 34244 | 1174 | let | 
| 35226 | 1175 | val (old_constrs, some_old_proj) = | 
| 35299 
4f4d5bf4ea08
proper distinction of code datatypes and abstypes
 haftmann parents: 
35226diff
changeset | 1176 | case these (Symtab.lookup ((the_types o the_exec) thy) tyco) | 
| 35226 | 1177 | of (_, (_, Constructors cos)) :: _ => (map fst cos, NONE) | 
| 36209 
566be5d48090
more appropriate representation of valid positions for abstractors; more accurate checking of abstype certificates
 haftmann parents: 
36202diff
changeset | 1178 | | (_, (_, Abstractor ((co, _), (proj, _)))) :: _ => ([co], SOME proj) | 
| 35226 | 1179 | | [] => ([], NONE) | 
| 1180 | val outdated_funs = case some_old_proj | |
| 35304 | 1181 | of NONE => old_constrs | 
| 35226 | 1182 | | SOME old_proj => Symtab.fold | 
| 36209 
566be5d48090
more appropriate representation of valid positions for abstractors; more accurate checking of abstype certificates
 haftmann parents: 
36202diff
changeset | 1183 | (fn (c, ((_, spec), _)) => | 
| 
566be5d48090
more appropriate representation of valid positions for abstractors; more accurate checking of abstype certificates
 haftmann parents: 
36202diff
changeset | 1184 | if member (op =) (the_list (associated_abstype spec)) tyco | 
| 35226 | 1185 | then insert (op =) c else I) | 
| 35304 | 1186 | ((the_functions o the_exec) thy) (old_proj :: old_constrs); | 
| 34244 | 1187 | fun drop_outdated_cases cases = fold Symtab.delete_safe | 
| 37438 | 1188 | (Symtab.fold (fn (c, ((_, (_, cos)), _)) => | 
| 35226 | 1189 | if exists (member (op =) old_constrs) cos | 
| 34244 | 1190 | then insert (op =) c else I) cases []) cases; | 
| 1191 | in | |
| 1192 | thy | |
| 35226 | 1193 | |> fold del_eqns outdated_funs | 
| 34244 | 1194 | |> map_exec_purge | 
| 35226 | 1195 | ((map_typs o Symtab.map_default (tyco, [])) (cons (serial (), vs_spec)) | 
| 34244 | 1196 | #> (map_cases o apfst) drop_outdated_cases) | 
| 1197 | end; | |
| 1198 | ||
| 35299 
4f4d5bf4ea08
proper distinction of code datatypes and abstypes
 haftmann parents: 
35226diff
changeset | 1199 | fun unoverload_const_typ thy (c, ty) = (AxClass.unoverload_const thy (c, ty), ty); | 
| 34244 | 1200 | |
| 35299 
4f4d5bf4ea08
proper distinction of code datatypes and abstypes
 haftmann parents: 
35226diff
changeset | 1201 | structure Datatype_Interpretation = | 
| 
4f4d5bf4ea08
proper distinction of code datatypes and abstypes
 haftmann parents: 
35226diff
changeset | 1202 | Interpretation(type T = string * serial val eq = eq_snd (op =) : T * T -> bool); | 
| 
4f4d5bf4ea08
proper distinction of code datatypes and abstypes
 haftmann parents: 
35226diff
changeset | 1203 | |
| 
4f4d5bf4ea08
proper distinction of code datatypes and abstypes
 haftmann parents: 
35226diff
changeset | 1204 | fun datatype_interpretation f = Datatype_Interpretation.interpretation | 
| 37448 
3bd4b3809bee
explicit type variable arguments for constructors
 haftmann parents: 
37438diff
changeset | 1205 | (fn (tyco, _) => fn thy => f (tyco, fst (get_type_spec thy tyco)) thy); | 
| 35226 | 1206 | |
| 1207 | fun add_datatype proto_constrs thy = | |
| 1208 | let | |
| 1209 | val constrs = map (unoverload_const_typ thy) proto_constrs; | |
| 1210 | val (tyco, (vs, cos)) = constrset_of_consts thy constrs; | |
| 1211 | in | |
| 1212 | thy | |
| 1213 | |> fold (del_eqns o fst) constrs | |
| 35299 
4f4d5bf4ea08
proper distinction of code datatypes and abstypes
 haftmann parents: 
35226diff
changeset | 1214 | |> register_type (tyco, (vs, Constructors cos)) | 
| 
4f4d5bf4ea08
proper distinction of code datatypes and abstypes
 haftmann parents: 
35226diff
changeset | 1215 | |> Datatype_Interpretation.data (tyco, serial ()) | 
| 35226 | 1216 | end; | 
| 1217 | ||
| 1218 | fun add_datatype_cmd raw_constrs thy = | |
| 1219 | add_datatype (map (read_bare_const thy) raw_constrs) thy; | |
| 1220 | ||
| 35299 
4f4d5bf4ea08
proper distinction of code datatypes and abstypes
 haftmann parents: 
35226diff
changeset | 1221 | structure Abstype_Interpretation = | 
| 
4f4d5bf4ea08
proper distinction of code datatypes and abstypes
 haftmann parents: 
35226diff
changeset | 1222 | Interpretation(type T = string * serial val eq = eq_snd (op =) : T * T -> bool); | 
| 
4f4d5bf4ea08
proper distinction of code datatypes and abstypes
 haftmann parents: 
35226diff
changeset | 1223 | |
| 
4f4d5bf4ea08
proper distinction of code datatypes and abstypes
 haftmann parents: 
35226diff
changeset | 1224 | fun abstype_interpretation f = Abstype_Interpretation.interpretation | 
| 
4f4d5bf4ea08
proper distinction of code datatypes and abstypes
 haftmann parents: 
35226diff
changeset | 1225 | (fn (tyco, _) => fn thy => f (tyco, get_abstype_spec thy tyco) thy); | 
| 
4f4d5bf4ea08
proper distinction of code datatypes and abstypes
 haftmann parents: 
35226diff
changeset | 1226 | |
| 36112 
7fa17a225852
user interface for abstract datatypes is an attribute, not a command
 haftmann parents: 
35845diff
changeset | 1227 | fun add_abstype proto_thm thy = | 
| 34244 | 1228 | let | 
| 36112 
7fa17a225852
user interface for abstract datatypes is an attribute, not a command
 haftmann parents: 
35845diff
changeset | 1229 | val (tyco, (vs, (abs_ty as (abs, ty), (rep, cert)))) = | 
| 36209 
566be5d48090
more appropriate representation of valid positions for abstractors; more accurate checking of abstype certificates
 haftmann parents: 
36202diff
changeset | 1230 | error_thm (check_abstype_cert thy) proto_thm; | 
| 35226 | 1231 | in | 
| 1232 | thy | |
| 36112 
7fa17a225852
user interface for abstract datatypes is an attribute, not a command
 haftmann parents: 
35845diff
changeset | 1233 | |> del_eqns abs | 
| 
7fa17a225852
user interface for abstract datatypes is an attribute, not a command
 haftmann parents: 
35845diff
changeset | 1234 | |> register_type (tyco, (vs, Abstractor (abs_ty, (rep, cert)))) | 
| 
7fa17a225852
user interface for abstract datatypes is an attribute, not a command
 haftmann parents: 
35845diff
changeset | 1235 | |> change_fun_spec false rep ((K o Proj) | 
| 
7fa17a225852
user interface for abstract datatypes is an attribute, not a command
 haftmann parents: 
35845diff
changeset | 1236 | (map_types Logic.varifyT_global (mk_proj tyco vs ty abs rep), tyco)) | 
| 
7fa17a225852
user interface for abstract datatypes is an attribute, not a command
 haftmann parents: 
35845diff
changeset | 1237 | |> Abstype_Interpretation.data (tyco, serial ()) | 
| 35226 | 1238 | end; | 
| 34244 | 1239 | |
| 35226 | 1240 | |
| 1241 | (** infrastructure **) | |
| 34244 | 1242 | |
| 32070 | 1243 | (* c.f. src/HOL/Tools/recfun_codegen.ML *) | 
| 1244 | ||
| 33522 | 1245 | structure Code_Target_Attr = Theory_Data | 
| 1246 | ( | |
| 31998 
2c7a24f74db9
code attributes use common underscore convention
 haftmann parents: 
31971diff
changeset | 1247 | type T = (string -> thm -> theory -> theory) option; | 
| 
2c7a24f74db9
code attributes use common underscore convention
 haftmann parents: 
31971diff
changeset | 1248 | val empty = NONE; | 
| 
2c7a24f74db9
code attributes use common underscore convention
 haftmann parents: 
31971diff
changeset | 1249 | val extend = I; | 
| 33522 | 1250 | fun merge (f1, f2) = if is_some f1 then f1 else f2; | 
| 31998 
2c7a24f74db9
code attributes use common underscore convention
 haftmann parents: 
31971diff
changeset | 1251 | ); | 
| 
2c7a24f74db9
code attributes use common underscore convention
 haftmann parents: 
31971diff
changeset | 1252 | |
| 32070 | 1253 | fun set_code_target_attr f = Code_Target_Attr.map (K (SOME f)); | 
| 1254 | ||
| 1255 | fun code_target_attr prefix thm thy = | |
| 1256 | let | |
| 1257 | val attr = the_default ((K o K) I) (Code_Target_Attr.get thy); | |
| 1258 | in thy |> add_warning_eqn thm |> attr prefix thm end; | |
| 35226 | 1259 | |
| 1260 | ||
| 32070 | 1261 | (* setup *) | 
| 31998 
2c7a24f74db9
code attributes use common underscore convention
 haftmann parents: 
31971diff
changeset | 1262 | |
| 31962 | 1263 | val _ = Context.>> (Context.map_theory | 
| 1264 | (let | |
| 1265 | 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 | 1266 | val code_attribute_parser = | 
| 
2c7a24f74db9
code attributes use common underscore convention
 haftmann parents: 
31971diff
changeset | 1267 | Args.del |-- Scan.succeed (mk_attribute del_eqn) | 
| 
2c7a24f74db9
code attributes use common underscore convention
 haftmann parents: 
31971diff
changeset | 1268 | || Args.$$$ "nbe" |-- Scan.succeed (mk_attribute add_nbe_eqn) | 
| 36112 
7fa17a225852
user interface for abstract datatypes is an attribute, not a command
 haftmann parents: 
35845diff
changeset | 1269 | || Args.$$$ "abstype" |-- Scan.succeed (mk_attribute add_abstype) | 
| 35226 | 1270 | || Args.$$$ "abstract" |-- Scan.succeed (mk_attribute add_abs_eqn) | 
| 31998 
2c7a24f74db9
code attributes use common underscore convention
 haftmann parents: 
31971diff
changeset | 1271 | || (Args.$$$ "target" |-- Args.colon |-- Args.name >> | 
| 32070 | 1272 | (mk_attribute o code_target_attr)) | 
| 31998 
2c7a24f74db9
code attributes use common underscore convention
 haftmann parents: 
31971diff
changeset | 1273 | || Scan.succeed (mk_attribute add_warning_eqn); | 
| 31962 | 1274 | in | 
| 35299 
4f4d5bf4ea08
proper distinction of code datatypes and abstypes
 haftmann parents: 
35226diff
changeset | 1275 | Datatype_Interpretation.init | 
| 31998 
2c7a24f74db9
code attributes use common underscore convention
 haftmann parents: 
31971diff
changeset | 1276 | #> Attrib.setup (Binding.name "code") (Scan.lift code_attribute_parser) | 
| 
2c7a24f74db9
code attributes use common underscore convention
 haftmann parents: 
31971diff
changeset | 1277 | "declare theorems for code generation" | 
| 31962 | 1278 | end)); | 
| 1279 | ||
| 24219 | 1280 | end; (*struct*) | 
| 1281 | ||
| 1282 | ||
| 35226 | 1283 | (* type-safe interfaces for data dependent on executable code *) | 
| 24219 | 1284 | |
| 34173 
458ced35abb8
reduced code generator cache to the baremost minimum
 haftmann parents: 
33977diff
changeset | 1285 | functor Code_Data(Data: CODE_DATA_ARGS): CODE_DATA = | 
| 24219 | 1286 | struct | 
| 1287 | ||
| 1288 | type T = Data.T; | |
| 1289 | exception Data of T; | |
| 1290 | fun dest (Data x) = x | |
| 1291 | ||
| 34173 
458ced35abb8
reduced code generator cache to the baremost minimum
 haftmann parents: 
33977diff
changeset | 1292 | val kind = Code.declare_data (Data Data.empty); | 
| 24219 | 1293 | |
| 1294 | val data_op = (kind, Data, dest); | |
| 1295 | ||
| 39397 | 1296 | fun change_yield (SOME thy) f = Code.change_yield_data data_op thy f | 
| 1297 | | change_yield NONE f = f Data.empty | |
| 1298 | ||
| 1299 | fun change some_thy f = snd (change_yield some_thy (pair () o f)); | |
| 24219 | 1300 | |
| 1301 | end; | |
| 1302 | ||
| 28143 | 1303 | structure Code : CODE = struct open Code; end; |