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