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