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