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