author | wenzelm |
Tue, 06 Jun 2023 21:02:37 +0200 | |
changeset 78140 | 90a43a9b3605 |
parent 78139 | bb85bda12b8e |
child 79232 | 99bc2dd45111 |
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_const: theory -> string -> string |
|
31962 | 14 |
val string_of_const: theory -> string -> string |
15 |
val args_number: theory -> string -> int |
|
31957 | 16 |
|
31156 | 17 |
(*constructor sets*) |
18 |
val constrset_of_consts: theory -> (string * typ) list |
|
40726
16dcfedc4eb7
keep type variable arguments of datatype constructors in bookkeeping
haftmann
parents:
40564
diff
changeset
|
19 |
-> string * ((string * sort) list * (string * ((string * sort) list * typ list)) list) |
31156 | 20 |
|
34874 | 21 |
(*code equations and certificates*) |
31156 | 22 |
val assert_eqn: theory -> thm * bool -> thm * bool |
66310
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
23 |
val assert_abs_eqn: theory -> string option -> thm -> thm * (string * string) |
34895 | 24 |
type cert |
34874 | 25 |
val constrain_cert: theory -> sort list -> cert -> cert |
49971
8b50286c36d3
close code theorems explicitly after preprocessing
haftmann
parents:
49904
diff
changeset
|
26 |
val conclude_cert: cert -> cert |
35226 | 27 |
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
|
28 |
val equations_of_cert: theory -> cert -> ((string * sort) list * typ) |
77702 | 29 |
* (((string option * term) list * (string option * term)) * (thm option * bool)) list option |
34895 | 30 |
val pretty_cert: theory -> cert -> Pretty.T list |
31156 | 31 |
|
31962 | 32 |
(*executable code*) |
66328 | 33 |
type constructors |
66330
dcb3e6bdc00a
one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents:
66329
diff
changeset
|
34 |
type abs_type |
dcb3e6bdc00a
one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents:
66329
diff
changeset
|
35 |
val type_interpretation: (string -> theory -> theory) -> theory -> theory |
dcb3e6bdc00a
one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents:
66329
diff
changeset
|
36 |
val datatype_interpretation: (string * constructors -> theory -> theory) -> theory -> theory |
dcb3e6bdc00a
one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents:
66329
diff
changeset
|
37 |
val abstype_interpretation: (string * abs_type -> theory -> theory) -> theory -> theory |
66251
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
38 |
val declare_datatype_global: (string * typ) list -> theory -> theory |
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
39 |
val declare_datatype_cmd: string list -> theory -> theory |
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
40 |
val declare_abstype: thm -> local_theory -> local_theory |
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
41 |
val declare_abstype_global: thm -> theory -> theory |
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
42 |
val declare_default_eqns: (thm * bool) list -> local_theory -> local_theory |
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
43 |
val declare_default_eqns_global: (thm * bool) list -> theory -> theory |
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
44 |
val declare_eqns: (thm * bool) list -> local_theory -> local_theory |
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
45 |
val declare_eqns_global: (thm * bool) list -> theory -> theory |
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
46 |
val add_eqn_global: thm * bool -> theory -> theory |
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
47 |
val del_eqn_global: thm -> theory -> theory |
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
48 |
val declare_abstract_eqn: thm -> local_theory -> local_theory |
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
49 |
val declare_abstract_eqn_global: thm -> theory -> theory |
66310
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
50 |
val declare_aborting_global: string -> theory -> theory |
66251
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
51 |
val declare_unimplemented_global: string -> theory -> theory |
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
52 |
val declare_case_global: thm -> theory -> theory |
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
53 |
val declare_undefined_global: string -> theory -> theory |
66328 | 54 |
val get_type: theory -> string -> constructors * bool |
35299
4f4d5bf4ea08
proper distinction of code datatypes and abstypes
haftmann
parents:
35226
diff
changeset
|
55 |
val get_type_of_constr_or_abstr: theory -> string -> (string * bool) option |
35226 | 56 |
val is_constr: theory -> string -> bool |
57 |
val is_abstr: theory -> string -> bool |
|
57429 | 58 |
val get_cert: Proof.context -> ((thm * bool) list -> (thm * bool) list option) list |
59 |
-> string -> cert |
|
66310
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
60 |
type case_schema |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
61 |
val get_case_schema: theory -> string -> case_schema option |
37438 | 62 |
val get_case_cong: theory -> string -> thm option |
66189
23917e861eaa
treat "undefined" constants internally as special form of case combinators
haftmann
parents:
66167
diff
changeset
|
63 |
val is_undefined: theory -> string -> bool |
24219 | 64 |
val print_codesetup: theory -> unit |
65 |
end; |
|
66 |
||
67 |
signature CODE_DATA_ARGS = |
|
68 |
sig |
|
69 |
type T |
|
70 |
val empty: T |
|
71 |
end; |
|
72 |
||
73 |
signature CODE_DATA = |
|
74 |
sig |
|
75 |
type T |
|
39397 | 76 |
val change: theory option -> (T -> T) -> T |
77 |
val change_yield: theory option -> (T -> 'a * T) -> 'a * T |
|
24219 | 78 |
end; |
79 |
||
80 |
signature PRIVATE_CODE = |
|
81 |
sig |
|
82 |
include CODE |
|
51368
2ea5c7c2d825
tuned signature -- prefer terminology of Scala and Axiom;
wenzelm
parents:
49971
diff
changeset
|
83 |
val declare_data: Any.T -> serial |
2ea5c7c2d825
tuned signature -- prefer terminology of Scala and Axiom;
wenzelm
parents:
49971
diff
changeset
|
84 |
val change_yield_data: serial * ('a -> Any.T) * (Any.T -> 'a) |
34251 | 85 |
-> theory -> ('a -> 'b * 'a) -> 'b * 'a |
24219 | 86 |
end; |
87 |
||
88 |
structure Code : PRIVATE_CODE = |
|
89 |
struct |
|
90 |
||
31962 | 91 |
(** auxiliary **) |
92 |
||
93 |
(* printing *) |
|
31156 | 94 |
|
39134
917b4b6ba3d2
turned show_sorts/show_types into proper configuration options;
wenzelm
parents:
39020
diff
changeset
|
95 |
fun string_of_typ thy = |
917b4b6ba3d2
turned show_sorts/show_types into proper configuration options;
wenzelm
parents:
39020
diff
changeset
|
96 |
Syntax.string_of_typ (Config.put show_sorts true (Syntax.init_pretty_global thy)); |
31962 | 97 |
|
42359 | 98 |
fun string_of_const thy c = |
42360 | 99 |
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
|
100 |
case Axclass.inst_of_param thy c of |
42359 | 101 |
SOME (c, tyco) => |
42360 | 102 |
Proof_Context.extern_const ctxt c ^ " " ^ enclose "[" "]" |
103 |
(Proof_Context.extern_type ctxt tyco) |
|
104 |
| NONE => Proof_Context.extern_const ctxt c |
|
42359 | 105 |
end; |
31156 | 106 |
|
31962 | 107 |
|
108 |
(* constants *) |
|
31156 | 109 |
|
49534
791e6fc53f73
more strict typscheme_equiv check: must fix variables of more specific type;
haftmann
parents:
49533
diff
changeset
|
110 |
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
|
111 |
|
791e6fc53f73
more strict typscheme_equiv check: must fix variables of more specific type;
haftmann
parents:
49533
diff
changeset
|
112 |
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
|
113 |
|
791e6fc53f73
more strict typscheme_equiv check: must fix variables of more specific type;
haftmann
parents:
49533
diff
changeset
|
114 |
fun devarify ty = |
791e6fc53f73
more strict typscheme_equiv check: must fix variables of more specific type;
haftmann
parents:
49533
diff
changeset
|
115 |
let |
75353 | 116 |
val tys = build (fold_atyps (fn TVar vi_sort => AList.update (op =) vi_sort) ty); |
49534
791e6fc53f73
more strict typscheme_equiv check: must fix variables of more specific type;
haftmann
parents:
49533
diff
changeset
|
117 |
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
|
118 |
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
|
119 |
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
|
120 |
|
791e6fc53f73
more strict typscheme_equiv check: must fix variables of more specific type;
haftmann
parents:
49533
diff
changeset
|
121 |
fun typscheme thy (c, ty) = |
791e6fc53f73
more strict typscheme_equiv check: must fix variables of more specific type;
haftmann
parents:
49533
diff
changeset
|
122 |
(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
|
123 |
|
791e6fc53f73
more strict typscheme_equiv check: must fix variables of more specific type;
haftmann
parents:
49533
diff
changeset
|
124 |
fun typscheme_equiv (ty1, ty2) = |
791e6fc53f73
more strict typscheme_equiv check: must fix variables of more specific type;
haftmann
parents:
49533
diff
changeset
|
125 |
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
|
126 |
|
31962 | 127 |
fun check_bare_const thy t = case try dest_Const t |
128 |
of SOME c_ty => c_ty |
|
129 |
| NONE => error ("Not a constant: " ^ Syntax.string_of_term_global thy t); |
|
31156 | 130 |
|
40362
82a066bff182
Code.check_const etc.: reject too specific types
haftmann
parents:
40316
diff
changeset
|
131 |
fun check_unoverload thy (c, ty) = |
82a066bff182
Code.check_const etc.: reject too specific types
haftmann
parents:
40316
diff
changeset
|
132 |
let |
51685
385ef6706252
more standard module name Axclass (according to file name);
wenzelm
parents:
51584
diff
changeset
|
133 |
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
|
134 |
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
|
135 |
in |
49534
791e6fc53f73
more strict typscheme_equiv check: must fix variables of more specific type;
haftmann
parents:
49533
diff
changeset
|
136 |
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
|
137 |
then c' |
e209da839ff4
added Logic.varify_types_global/unvarify_types_global, which avoids somewhat expensive Term.map_types;
wenzelm
parents:
45211
diff
changeset
|
138 |
else |
e209da839ff4
added Logic.varify_types_global/unvarify_types_global, which avoids somewhat expensive Term.map_types;
wenzelm
parents:
45211
diff
changeset
|
139 |
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
|
140 |
"\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
|
141 |
"\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
|
142 |
string_of_typ thy ty_decl) |
66167 | 143 |
end; |
40362
82a066bff182
Code.check_const etc.: reject too specific types
haftmann
parents:
40316
diff
changeset
|
144 |
|
82a066bff182
Code.check_const etc.: reject too specific types
haftmann
parents:
40316
diff
changeset
|
145 |
fun check_const thy = check_unoverload thy o check_bare_const thy; |
31962 | 146 |
|
147 |
fun read_bare_const thy = check_bare_const thy o Syntax.read_term_global thy; |
|
148 |
||
40362
82a066bff182
Code.check_const etc.: reject too specific types
haftmann
parents:
40316
diff
changeset
|
149 |
fun read_const thy = check_unoverload thy o read_bare_const thy; |
31156 | 150 |
|
32872
019201eb7e07
variables in type schemes must be renamed simultaneously with variables in equations
haftmann
parents:
32738
diff
changeset
|
151 |
|
66310
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
152 |
(** executable specifications **) |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
153 |
|
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
154 |
(* types *) |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
155 |
|
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
156 |
datatype type_spec = Constructors of { |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
157 |
constructors: (string * ((string * sort) list * typ list)) list, |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
158 |
case_combinators: string list} |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
159 |
| Abstractor of { |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
160 |
abs_rep: thm, |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
161 |
abstractor: string * ((string * sort) list * typ), |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
162 |
projection: string, |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
163 |
more_abstract_functions: string list}; |
31962 | 164 |
|
66310
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
165 |
fun concrete_constructors_of (Constructors {constructors, ...}) = |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
166 |
constructors |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
167 |
| concrete_constructors_of _ = |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
168 |
[]; |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
169 |
|
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
170 |
fun constructors_of (Constructors {constructors, ...}) = |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
171 |
(constructors, false) |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
172 |
| constructors_of (Abstractor {abstractor = (co, (vs, ty)), ...}) = |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
173 |
([(co, (vs, [ty]))], true); |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
174 |
|
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
175 |
fun case_combinators_of (Constructors {case_combinators, ...}) = |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
176 |
case_combinators |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
177 |
| case_combinators_of (Abstractor _) = |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
178 |
[]; |
35226 | 179 |
|
66310
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
180 |
fun add_case_combinator c (vs, Constructors {constructors, case_combinators}) = |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
181 |
(vs, Constructors {constructors = constructors, |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
182 |
case_combinators = insert (op =) c case_combinators}); |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
183 |
|
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
184 |
fun projection_of (Constructors _) = |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
185 |
NONE |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
186 |
| projection_of (Abstractor {projection, ...}) = |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
187 |
SOME projection; |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
188 |
|
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
189 |
fun abstract_functions_of (Constructors _) = |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
190 |
[] |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
191 |
| abstract_functions_of (Abstractor {more_abstract_functions, projection, ...}) = |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
192 |
projection :: more_abstract_functions; |
31962 | 193 |
|
66310
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
194 |
fun add_abstract_function c (vs, Abstractor {abs_rep, abstractor, projection, more_abstract_functions}) = |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
195 |
(vs, Abstractor {abs_rep = abs_rep, abstractor = abstractor, projection = projection, |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
196 |
more_abstract_functions = insert (op =) c more_abstract_functions}); |
35226 | 197 |
|
66310
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
198 |
fun join_same_types' (Constructors {constructors, case_combinators = case_combinators1}, |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
199 |
Constructors {case_combinators = case_combinators2, ...}) = |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
200 |
Constructors {constructors = constructors, |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
201 |
case_combinators = merge (op =) (case_combinators1, case_combinators2)} |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
202 |
| join_same_types' (Abstractor {abs_rep, abstractor, projection, more_abstract_functions = more_abstract_functions1}, |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
203 |
Abstractor {more_abstract_functions = more_abstract_functions2, ...}) = |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
204 |
Abstractor {abs_rep = abs_rep, abstractor = abstractor, projection = projection, |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
205 |
more_abstract_functions = merge (op =) (more_abstract_functions1, more_abstract_functions2)}; |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
206 |
|
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
207 |
fun join_same_types ((vs, spec1), (_, spec2)) = (vs, join_same_types' (spec1, spec2)); |
35226 | 208 |
|
66125 | 209 |
|
35226 | 210 |
(* functions *) |
31962 | 211 |
|
66310
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
212 |
datatype fun_spec = |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
213 |
Eqns of bool * (thm * bool) list |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
214 |
| Proj of term * (string * string) |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
215 |
| Abstr of thm * (string * string); |
31962 | 216 |
|
66310
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
217 |
val unimplemented = Eqns (true, []); |
31962 | 218 |
|
66310
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
219 |
fun is_unimplemented (Eqns (true, [])) = true |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
220 |
| is_unimplemented _ = false; |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
221 |
|
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
222 |
fun is_default (Eqns (true, _)) = true |
35226 | 223 |
| is_default _ = false; |
224 |
||
66310
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
225 |
val aborting = Eqns (false, []); |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
226 |
|
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
227 |
fun associated_abstype (Proj (_, tyco_abs)) = SOME tyco_abs |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
228 |
| associated_abstype (Abstr (_, tyco_abs)) = SOME tyco_abs |
35226 | 229 |
| associated_abstype _ = NONE; |
31962 | 230 |
|
231 |
||
66189
23917e861eaa
treat "undefined" constants internally as special form of case combinators
haftmann
parents:
66167
diff
changeset
|
232 |
(* cases *) |
23917e861eaa
treat "undefined" constants internally as special form of case combinators
haftmann
parents:
66167
diff
changeset
|
233 |
|
75399
cdf84288d93c
pass constructor arity as part of case certficiate
haftmann
parents:
75353
diff
changeset
|
234 |
type case_schema = int * (int * (string * int) option list); |
66310
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
235 |
|
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
236 |
datatype case_spec = |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
237 |
No_Case |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
238 |
| Case of {schema: case_schema, tycos: string list, cong: thm} |
66189
23917e861eaa
treat "undefined" constants internally as special form of case combinators
haftmann
parents:
66167
diff
changeset
|
239 |
| Undefined; |
23917e861eaa
treat "undefined" constants internally as special form of case combinators
haftmann
parents:
66167
diff
changeset
|
240 |
|
75399
cdf84288d93c
pass constructor arity as part of case certficiate
haftmann
parents:
75353
diff
changeset
|
241 |
fun associated_datatypes (Case {tycos, schema = (_, (_, raw_cos)), ...}) = (tycos, map fst (map_filter I raw_cos)) |
66310
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
242 |
| associated_datatypes _ = ([], []); |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
243 |
|
66189
23917e861eaa
treat "undefined" constants internally as special form of case combinators
haftmann
parents:
66167
diff
changeset
|
244 |
|
66310
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
245 |
(** background theory data store **) |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
246 |
|
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
247 |
(* historized declaration data *) |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
248 |
|
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
249 |
structure History = |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
250 |
struct |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
251 |
|
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
252 |
type 'a T = { |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
253 |
entry: 'a, |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
254 |
suppressed: bool, (*incompatible entries are merely suppressed after theory merge but sustain*) |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
255 |
history: serial list (*explicit trace of declaration history supports non-monotonic declarations*) |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
256 |
} Symtab.table; |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
257 |
|
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
258 |
fun some_entry (SOME {suppressed = false, entry, ...}) = SOME entry |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
259 |
| some_entry _ = NONE; |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
260 |
|
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
261 |
fun lookup table = |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
262 |
Symtab.lookup table #> some_entry; |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
263 |
|
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
264 |
fun register key entry table = |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
265 |
if is_some (Symtab.lookup table key) |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
266 |
then Symtab.map_entry key |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
267 |
(fn {history, ...} => {entry = entry, suppressed = false, history = serial () :: history}) table |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
268 |
else Symtab.update (key, {entry = entry, suppressed = false, history = [serial ()]}) table; |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
269 |
|
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
270 |
fun modify_entry key f = Symtab.map_entry key |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
271 |
(fn {entry, suppressed, history} => {entry = f entry, suppressed = suppressed, history = history}); |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
272 |
|
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
273 |
fun all table = Symtab.dest table |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
274 |
|> map_filter (fn (key, {entry, suppressed = false, ...}) => SOME (key, entry) | _ => NONE); |
31962 | 275 |
|
66310
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
276 |
local |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
277 |
|
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
278 |
fun merge_history join_same |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
279 |
({entry = entry1, history = history1, ...}, {entry = entry2, history = history2, ...}) = |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
280 |
let |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
281 |
val history = merge (op =) (history1, history2); |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
282 |
val entry = if hd history1 = hd history2 then join_same (entry1, entry2) |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
283 |
else if hd history = hd history1 then entry1 else entry2; |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
284 |
in {entry = entry, suppressed = false, history = history} end; |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
285 |
|
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
286 |
in |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
287 |
|
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
288 |
fun join join_same tables = Symtab.join (K (merge_history join_same)) tables; |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
289 |
|
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
290 |
fun suppress key = Symtab.map_entry key |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
291 |
(fn {entry, history, ...} => {entry = entry, suppressed = true, history = history}); |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
292 |
|
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
293 |
fun suppress_except f = Symtab.map (fn key => fn {entry, suppressed, history} => |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
294 |
{entry = entry, suppressed = suppressed orelse (not o f) (key, entry), history = history}); |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
295 |
|
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
296 |
end; |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
297 |
|
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
298 |
end; |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
299 |
|
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
300 |
datatype specs = Specs of { |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
301 |
types: ((string * sort) list * type_spec) History.T, |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
302 |
pending_eqns: (thm * bool) list Symtab.table, |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
303 |
functions: fun_spec History.T, |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
304 |
cases: case_spec History.T |
31962 | 305 |
}; |
306 |
||
66310
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
307 |
fun types_of (Specs {types, ...}) = types; |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
308 |
fun pending_eqns_of (Specs {pending_eqns, ...}) = pending_eqns; |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
309 |
fun functions_of (Specs {functions, ...}) = functions; |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
310 |
fun cases_of (Specs {cases, ...}) = cases; |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
311 |
|
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
312 |
fun make_specs (types, ((pending_eqns, functions), cases)) = |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
313 |
Specs {types = types, pending_eqns = pending_eqns, |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
314 |
functions = functions, cases = cases}; |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
315 |
|
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
316 |
val empty_specs = |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
317 |
make_specs (Symtab.empty, ((Symtab.empty, Symtab.empty), Symtab.empty)); |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
318 |
|
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
319 |
fun map_specs f (Specs {types = types, pending_eqns = pending_eqns, |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
320 |
functions = functions, cases = cases}) = |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
321 |
make_specs (f (types, ((pending_eqns, functions), cases))); |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
322 |
|
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
323 |
fun merge_specs (Specs {types = types1, pending_eqns = _, |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
324 |
functions = functions1, cases = cases1}, |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
325 |
Specs {types = types2, pending_eqns = _, |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
326 |
functions = functions2, cases = cases2}) = |
31156 | 327 |
let |
66310
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
328 |
val types = History.join join_same_types (types1, types2); |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
329 |
val all_types = map (snd o snd) (History.all types); |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
330 |
fun check_abstype (c, fun_spec) = case associated_abstype fun_spec of |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
331 |
NONE => true |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
332 |
| SOME (tyco, abs) => (case History.lookup types tyco of |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
333 |
NONE => false |
70358 | 334 |
| SOME (_, Constructors _) => false |
66310
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
335 |
| SOME (_, Abstractor {abstractor = (abs', _), projection, more_abstract_functions, ...}) => |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
336 |
abs = abs' andalso (c = projection orelse member (op =) more_abstract_functions c)); |
66324 | 337 |
fun check_datatypes (_, case_spec) = |
31962 | 338 |
let |
66310
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
339 |
val (tycos, required_constructors) = associated_datatypes case_spec; |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
340 |
val allowed_constructors = |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
341 |
tycos |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
342 |
|> maps (these o Option.map (concrete_constructors_of o snd) o History.lookup types) |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
343 |
|> map fst; |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
344 |
in subset (op =) (required_constructors, allowed_constructors) end; |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
345 |
val all_constructors = |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
346 |
maps (fst o constructors_of) all_types; |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
347 |
val functions = History.join fst (functions1, functions2) |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
348 |
|> fold (History.suppress o fst) all_constructors |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
349 |
|> History.suppress_except check_abstype; |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
350 |
val cases = History.join fst (cases1, cases2) |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
351 |
|> History.suppress_except check_datatypes; |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
352 |
in make_specs (types, ((Symtab.empty, functions), cases)) end; |
31962 | 353 |
|
66310
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
354 |
val map_types = map_specs o apfst; |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
355 |
val map_pending_eqns = map_specs o apsnd o apfst o apfst; |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
356 |
val map_functions = map_specs o apsnd o apfst o apsnd; |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
357 |
val map_cases = map_specs o apsnd o apsnd; |
31962 | 358 |
|
359 |
||
360 |
(* data slots dependent on executable code *) |
|
361 |
||
362 |
(*private copy avoids potential conflict of table exceptions*) |
|
31971
8c1b845ed105
renamed functor TableFun to Table, and GraphFun to Graph;
wenzelm
parents:
31962
diff
changeset
|
363 |
structure Datatab = Table(type key = int val ord = int_ord); |
31962 | 364 |
|
365 |
local |
|
366 |
||
66310
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
367 |
type kind = {empty: Any.T}; |
31962 | 368 |
|
43684 | 369 |
val kinds = Synchronized.var "Code_Data" (Datatab.empty: kind Datatab.table); |
31962 | 370 |
|
43684 | 371 |
fun invoke f k = |
372 |
(case Datatab.lookup (Synchronized.value kinds) k of |
|
373 |
SOME kind => f kind |
|
374 |
| NONE => raise Fail "Invalid code data identifier"); |
|
31962 | 375 |
|
376 |
in |
|
377 |
||
34173
458ced35abb8
reduced code generator cache to the baremost minimum
haftmann
parents:
33977
diff
changeset
|
378 |
fun declare_data empty = |
31962 | 379 |
let |
380 |
val k = serial (); |
|
66310
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
381 |
val kind = {empty = empty}; |
43684 | 382 |
val _ = Synchronized.change kinds (Datatab.update (k, kind)); |
31962 | 383 |
in k end; |
384 |
||
385 |
fun invoke_init k = invoke (fn kind => #empty kind) k; |
|
386 |
||
387 |
end; (*local*) |
|
388 |
||
389 |
||
66310
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
390 |
(* global theory store *) |
31962 | 391 |
|
392 |
local |
|
393 |
||
51368
2ea5c7c2d825
tuned signature -- prefer terminology of Scala and Axiom;
wenzelm
parents:
49971
diff
changeset
|
394 |
type data = Any.T Datatab.table; |
72057 | 395 |
|
78053
64f81d011a90
clarified signature: avoid convoluted operations;
wenzelm
parents:
77928
diff
changeset
|
396 |
fun make_dataref () = |
64f81d011a90
clarified signature: avoid convoluted operations;
wenzelm
parents:
77928
diff
changeset
|
397 |
Synchronized.var "code data" (NONE : (data * Context.theory_id) option); |
31962 | 398 |
|
34244 | 399 |
structure Code_Data = Theory_Data |
31962 | 400 |
( |
72057 | 401 |
type T = specs * (string * (data * Context.theory_id) option Synchronized.var); |
78053
64f81d011a90
clarified signature: avoid convoluted operations;
wenzelm
parents:
77928
diff
changeset
|
402 |
val empty = |
64f81d011a90
clarified signature: avoid convoluted operations;
wenzelm
parents:
77928
diff
changeset
|
403 |
(empty_specs, (Context.theory_long_name (Context.the_global_context ()), make_dataref ())); |
72057 | 404 |
fun merge ((specs1, dataref), (specs2, _)) = |
405 |
(merge_specs (specs1, specs2), dataref); |
|
31962 | 406 |
); |
407 |
||
72057 | 408 |
fun init_dataref thy = |
78053
64f81d011a90
clarified signature: avoid convoluted operations;
wenzelm
parents:
77928
diff
changeset
|
409 |
let val thy_name = Context.theory_long_name thy in |
64f81d011a90
clarified signature: avoid convoluted operations;
wenzelm
parents:
77928
diff
changeset
|
410 |
if #1 (#2 (Code_Data.get thy)) = thy_name then NONE |
64f81d011a90
clarified signature: avoid convoluted operations;
wenzelm
parents:
77928
diff
changeset
|
411 |
else SOME ((Code_Data.map o apsnd) (K (thy_name, make_dataref ())) thy) |
64f81d011a90
clarified signature: avoid convoluted operations;
wenzelm
parents:
77928
diff
changeset
|
412 |
end; |
72057 | 413 |
|
31962 | 414 |
in |
415 |
||
72057 | 416 |
val _ = Theory.setup (Theory.at_begin init_dataref); |
417 |
||
35226 | 418 |
|
66310
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
419 |
(* access to executable specifications *) |
31962 | 420 |
|
66310
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
421 |
val specs_of : theory -> specs = fst o Code_Data.get; |
31962 | 422 |
|
72057 | 423 |
fun modify_specs f thy = |
78053
64f81d011a90
clarified signature: avoid convoluted operations;
wenzelm
parents:
77928
diff
changeset
|
424 |
let val thy_name = Context.theory_long_name thy |
64f81d011a90
clarified signature: avoid convoluted operations;
wenzelm
parents:
77928
diff
changeset
|
425 |
in Code_Data.map (fn (specs, _) => (f specs, (thy_name, make_dataref ()))) thy end; |
31962 | 426 |
|
427 |
||
66310
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
428 |
(* access to data dependent on executable specifications *) |
31962 | 429 |
|
34244 | 430 |
fun change_yield_data (kind, mk, dest) theory f = |
31962 | 431 |
let |
72057 | 432 |
val dataref = #2 (#2 (Code_Data.get theory)); |
67646 | 433 |
val (datatab, thy_id) = case Synchronized.value dataref |
434 |
of SOME (datatab, thy_id) => |
|
435 |
if Context.eq_thy_id (Context.theory_id theory, thy_id) |
|
436 |
then (datatab, thy_id) |
|
437 |
else (Datatab.empty, Context.theory_id theory) |
|
438 |
| NONE => (Datatab.empty, Context.theory_id theory) |
|
34244 | 439 |
val data = case Datatab.lookup datatab kind |
440 |
of SOME data => data |
|
441 |
| NONE => invoke_init kind; |
|
40758
4f2c3e842ef8
consider sort constraints for datatype constructors when constructing the empty equation certificate;
haftmann
parents:
40726
diff
changeset
|
442 |
val result as (_, data') = f (dest data); |
34244 | 443 |
val _ = Synchronized.change dataref |
67646 | 444 |
((K o SOME) (Datatab.update (kind, mk data') datatab, thy_id)); |
34244 | 445 |
in result end; |
31962 | 446 |
|
447 |
end; (*local*) |
|
448 |
||
449 |
||
66310
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
450 |
(* pending function equations *) |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
451 |
|
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
452 |
(* Ideally, *all* equations implementing a functions would be treated as |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
453 |
*one* atomic declaration; unfortunately, we cannot implement this: |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
454 |
the too-well-established declaration interface are Isar attributes |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
455 |
which operate on *one* single theorem. Hence we treat such Isar |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
456 |
declarations as "pending" and historize them as proper declarations |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
457 |
at the end of each theory. *) |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
458 |
|
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
459 |
fun modify_pending_eqns c f specs = |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
460 |
let |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
461 |
val existing_eqns = case History.lookup (functions_of specs) c of |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
462 |
SOME (Eqns (false, eqns)) => eqns |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
463 |
| _ => []; |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
464 |
in |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
465 |
specs |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
466 |
|> map_pending_eqns (Symtab.map_default (c, existing_eqns) f) |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
467 |
end; |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
468 |
|
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
469 |
fun register_fun_spec c spec = |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
470 |
map_pending_eqns (Symtab.delete_safe c) |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
471 |
#> map_functions (History.register c spec); |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
472 |
|
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
473 |
fun lookup_fun_spec specs c = |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
474 |
case Symtab.lookup (pending_eqns_of specs) c of |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
475 |
SOME eqns => Eqns (false, eqns) |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
476 |
| NONE => (case History.lookup (functions_of specs) c of |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
477 |
SOME spec => spec |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
478 |
| NONE => unimplemented); |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
479 |
|
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
480 |
fun lookup_proper_fun_spec specs c = |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
481 |
let |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
482 |
val spec = lookup_fun_spec specs c |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
483 |
in |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
484 |
if is_unimplemented spec then NONE else SOME spec |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
485 |
end; |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
486 |
|
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
487 |
fun all_fun_specs specs = |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
488 |
map_filter (fn c => Option.map (pair c) (lookup_proper_fun_spec specs c)) |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
489 |
(union (op =) |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
490 |
((Symtab.keys o pending_eqns_of) specs) |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
491 |
((Symtab.keys o functions_of) specs)); |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
492 |
|
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
493 |
fun historize_pending_fun_specs thy = |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
494 |
let |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
495 |
val pending_eqns = (pending_eqns_of o specs_of) thy; |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
496 |
in if Symtab.is_empty pending_eqns |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
497 |
then |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
498 |
NONE |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
499 |
else |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
500 |
thy |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
501 |
|> modify_specs (map_functions |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
502 |
(Symtab.fold (fn (c, eqs) => History.register c (Eqns (false, eqs))) pending_eqns) |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
503 |
#> map_pending_eqns (K Symtab.empty)) |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
504 |
|> SOME |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
505 |
end; |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
506 |
|
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
507 |
val _ = Theory.setup (Theory.at_end historize_pending_fun_specs); |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
508 |
|
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
509 |
|
31962 | 510 |
(** foundation **) |
511 |
||
66128 | 512 |
(* types *) |
31156 | 513 |
|
35226 | 514 |
fun no_constr thy s (c, ty) = error ("Not a datatype constructor:\n" ^ string_of_const thy c |
515 |
^ " :: " ^ string_of_typ thy ty ^ "\n" ^ enclose "(" ")" s); |
|
516 |
||
45987 | 517 |
fun analyze_constructor thy (c, ty) = |
31156 | 518 |
let |
59621
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents:
59458
diff
changeset
|
519 |
val _ = Thm.global_cterm_of thy (Const (c, ty)); |
54603
89d5b69e5a08
prefer name-normalizing devarify over unvarifyT whenever appropriate
haftmann
parents:
53171
diff
changeset
|
520 |
val ty_decl = devarify (const_typ thy c); |
31156 | 521 |
fun last_typ c_ty ty = |
522 |
let |
|
33531 | 523 |
val tfrees = Term.add_tfreesT ty []; |
40844 | 524 |
val (tyco, vs) = (apsnd o map) dest_TFree (dest_Type (body_type ty)) |
35226 | 525 |
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
|
526 |
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
|
527 |
val _ = |
e209da839ff4
added Logic.varify_types_global/unvarify_types_global, which avoids somewhat expensive Term.map_types;
wenzelm
parents:
45211
diff
changeset
|
528 |
if has_duplicates (eq_fst (op =)) vs |
35226 | 529 |
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
|
530 |
val _ = |
e209da839ff4
added Logic.varify_types_global/unvarify_types_global, which avoids somewhat expensive Term.map_types;
wenzelm
parents:
45211
diff
changeset
|
531 |
if length tfrees <> length vs |
35226 | 532 |
then no_constr thy "type variables missing in datatype" c_ty else (); |
31156 | 533 |
in (tyco, vs) end; |
35226 | 534 |
val (tyco, _) = last_typ (c, ty) ty_decl; |
535 |
val (_, vs) = last_typ (c, ty) ty; |
|
536 |
in ((tyco, map snd vs), (c, (map fst vs, ty))) end; |
|
537 |
||
49904
2df2786ac7b7
no sort constraints on datatype constructors in internal bookkeeping
haftmann
parents:
49760
diff
changeset
|
538 |
fun constrset_of_consts thy consts = |
35226 | 539 |
let |
51685
385ef6706252
more standard module name Axclass (according to file name);
wenzelm
parents:
51584
diff
changeset
|
540 |
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
|
541 |
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
|
542 |
val raw_constructors = map (analyze_constructor thy) consts; |
2df2786ac7b7
no sort constraints on datatype constructors in internal bookkeeping
haftmann
parents:
49760
diff
changeset
|
543 |
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
|
544 |
of [tyco] => tyco |
2df2786ac7b7
no sort constraints on datatype constructors in internal bookkeeping
haftmann
parents:
49760
diff
changeset
|
545 |
| [] => error "Empty constructor set" |
2df2786ac7b7
no sort constraints on datatype constructors in internal bookkeeping
haftmann
parents:
49760
diff
changeset
|
546 |
| 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
|
547 |
val vs = Name.invent Name.context Name.aT (Sign.arity_number thy tyco); |
31156 | 548 |
fun inst vs' (c, (vs, ty)) = |
549 |
let |
|
550 |
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
|
551 |
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
|
552 |
val (vs'', ty'') = typscheme thy (c, ty'); |
2df2786ac7b7
no sort constraints on datatype constructors in internal bookkeeping
haftmann
parents:
49760
diff
changeset
|
553 |
in (c, (vs'', binder_types ty'')) end; |
2df2786ac7b7
no sort constraints on datatype constructors in internal bookkeeping
haftmann
parents:
49760
diff
changeset
|
554 |
val constructors = map (inst vs o snd) raw_constructors; |
2df2786ac7b7
no sort constraints on datatype constructors in internal bookkeeping
haftmann
parents:
49760
diff
changeset
|
555 |
in (tyco, (map (rpair []) vs, constructors)) end; |
31156 | 556 |
|
66310
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
557 |
fun lookup_vs_type_spec thy = History.lookup ((types_of o specs_of) thy); |
31962 | 558 |
|
66328 | 559 |
type constructors = |
560 |
(string * sort) list * (string * ((string * sort) list * typ list)) list; |
|
561 |
||
66310
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
562 |
fun get_type thy tyco = case lookup_vs_type_spec thy tyco |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
563 |
of SOME (vs, type_spec) => apfst (pair vs) (constructors_of type_spec) |
45987 | 564 |
| NONE => Sign.arity_number thy tyco |
43329
84472e198515
tuned signature: Name.invent and Name.invent_names;
wenzelm
parents:
43326
diff
changeset
|
565 |
|> Name.invent Name.context Name.aT |
35226 | 566 |
|> map (rpair []) |
567 |
|> rpair [] |
|
568 |
|> rpair false; |
|
569 |
||
66328 | 570 |
type abs_type = |
571 |
(string * sort) list * {abs_rep: thm, abstractor: string * ((string * sort) list * typ), projection: string}; |
|
572 |
||
66310
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
573 |
fun get_abstype_spec thy tyco = case lookup_vs_type_spec thy tyco of |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
574 |
SOME (vs, Abstractor {abs_rep, abstractor, projection, ...}) => |
78056 | 575 |
(vs, {abs_rep = Thm.transfer thy abs_rep, abstractor = abstractor, projection = projection}) |
36122 | 576 |
| _ => error ("Not an abstract type: " ^ tyco); |
66167 | 577 |
|
35299
4f4d5bf4ea08
proper distinction of code datatypes and abstypes
haftmann
parents:
35226
diff
changeset
|
578 |
fun get_type_of_constr_or_abstr thy c = |
40844 | 579 |
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
|
580 |
of Type (tyco, _) => let val ((_, cos), abstract) = get_type thy tyco |
35226 | 581 |
in if member (op =) (map fst cos) c then SOME (tyco, abstract) else NONE end |
31962 | 582 |
| _ => NONE; |
583 |
||
35299
4f4d5bf4ea08
proper distinction of code datatypes and abstypes
haftmann
parents:
35226
diff
changeset
|
584 |
fun is_constr thy c = case get_type_of_constr_or_abstr thy c |
35226 | 585 |
of SOME (_, false) => true |
586 |
| _ => false; |
|
587 |
||
35299
4f4d5bf4ea08
proper distinction of code datatypes and abstypes
haftmann
parents:
35226
diff
changeset
|
588 |
fun is_abstr thy c = case get_type_of_constr_or_abstr thy c |
35226 | 589 |
of SOME (_, true) => true |
590 |
| _ => false; |
|
31962 | 591 |
|
31156 | 592 |
|
34874 | 593 |
(* bare code equations *) |
31156 | 594 |
|
35226 | 595 |
(* convention for variables: |
596 |
?x ?'a for free-floating theorems (e.g. in the data store) |
|
597 |
?x 'a for certificates |
|
598 |
x 'a for final representation of equations |
|
599 |
*) |
|
600 |
||
31156 | 601 |
exception BAD_THM of string; |
63242
9986559617ee
clear distinction between different situations concerning strictness of code equations
haftmann
parents:
63241
diff
changeset
|
602 |
|
9986559617ee
clear distinction between different situations concerning strictness of code equations
haftmann
parents:
63241
diff
changeset
|
603 |
datatype strictness = Silent | Liberal | Strict |
9986559617ee
clear distinction between different situations concerning strictness of code equations
haftmann
parents:
63241
diff
changeset
|
604 |
|
9986559617ee
clear distinction between different situations concerning strictness of code equations
haftmann
parents:
63241
diff
changeset
|
605 |
fun handle_strictness thm_of f strictness thy x = SOME (f x) |
9986559617ee
clear distinction between different situations concerning strictness of code equations
haftmann
parents:
63241
diff
changeset
|
606 |
handle BAD_THM msg => case strictness of |
9986559617ee
clear distinction between different situations concerning strictness of code equations
haftmann
parents:
63241
diff
changeset
|
607 |
Silent => NONE |
9986559617ee
clear distinction between different situations concerning strictness of code equations
haftmann
parents:
63241
diff
changeset
|
608 |
| Liberal => (warning (msg ^ ", in theorem:\n" ^ Thm.string_of_thm_global thy (thm_of x)); NONE) |
9986559617ee
clear distinction between different situations concerning strictness of code equations
haftmann
parents:
63241
diff
changeset
|
609 |
| Strict => error (msg ^ ", in theorem:\n" ^ Thm.string_of_thm_global thy (thm_of x)); |
31156 | 610 |
|
611 |
fun is_linear thm = |
|
63242
9986559617ee
clear distinction between different situations concerning strictness of code equations
haftmann
parents:
63241
diff
changeset
|
612 |
let |
9986559617ee
clear distinction between different situations concerning strictness of code equations
haftmann
parents:
63241
diff
changeset
|
613 |
val (_, args) = (strip_comb o fst o Logic.dest_equals o Thm.plain_prop_of) thm |
9986559617ee
clear distinction between different situations concerning strictness of code equations
haftmann
parents:
63241
diff
changeset
|
614 |
in |
9986559617ee
clear distinction between different situations concerning strictness of code equations
haftmann
parents:
63241
diff
changeset
|
615 |
not (has_duplicates (op =) ((fold o fold_aterms) |
9986559617ee
clear distinction between different situations concerning strictness of code equations
haftmann
parents:
63241
diff
changeset
|
616 |
(fn Var (v, _) => cons v | _ => I) args [])) |
9986559617ee
clear distinction between different situations concerning strictness of code equations
haftmann
parents:
63241
diff
changeset
|
617 |
end; |
31156 | 618 |
|
36209
566be5d48090
more appropriate representation of valid positions for abstractors; more accurate checking of abstype certificates
haftmann
parents:
36202
diff
changeset
|
619 |
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
|
620 |
let |
54604
1512fa5fe531
prefer sort-stripping const_typ over Sign.the_const_type whenever appropriate
haftmann
parents:
54603
diff
changeset
|
621 |
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
|
622 |
in if typscheme_equiv (ty_decl, ty) then () |
78139 | 623 |
else raise BAD_THM ("Type\n" ^ string_of_typ thy ty |
36209
566be5d48090
more appropriate representation of valid positions for abstractors; more accurate checking of abstype certificates
haftmann
parents:
36202
diff
changeset
|
624 |
^ "\nof constant " ^ quote c |
40362
82a066bff182
Code.check_const etc.: reject too specific types
haftmann
parents:
40316
diff
changeset
|
625 |
^ "\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
|
626 |
^ string_of_typ thy ty_decl) |
66167 | 627 |
end; |
36209
566be5d48090
more appropriate representation of valid positions for abstractors; more accurate checking of abstype certificates
haftmann
parents:
36202
diff
changeset
|
628 |
|
66310
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
629 |
fun check_eqn thy {allow_nonlinear, allow_consts, allow_pats} thm (lhs, rhs) = |
31156 | 630 |
let |
631 |
fun vars_of t = fold_aterms (fn Var (v, _) => insert (op =) v |
|
78139 | 632 |
| Free _ => raise BAD_THM "Illegal free variable" |
31156 | 633 |
| _ => I) t []; |
634 |
fun tvars_of t = fold_term_types (fn _ => |
|
635 |
fold_atyps (fn TVar (v, _) => insert (op =) v |
|
78139 | 636 |
| TFree _ => raise BAD_THM "Illegal free type variable")) t []; |
31156 | 637 |
val lhs_vs = vars_of lhs; |
638 |
val rhs_vs = vars_of rhs; |
|
639 |
val lhs_tvs = tvars_of lhs; |
|
640 |
val rhs_tvs = tvars_of rhs; |
|
641 |
val _ = if null (subtract (op =) lhs_vs rhs_vs) |
|
642 |
then () |
|
78139 | 643 |
else raise BAD_THM "Free variables on right hand side of equation"; |
31156 | 644 |
val _ = if null (subtract (op =) lhs_tvs rhs_tvs) |
645 |
then () |
|
78139 | 646 |
else raise BAD_THM "Free type variables on right hand side of equation"; |
34894 | 647 |
val (head, args) = strip_comb lhs; |
31156 | 648 |
val (c, ty) = case head |
51685
385ef6706252
more standard module name Axclass (according to file name);
wenzelm
parents:
51584
diff
changeset
|
649 |
of Const (c_ty as (_, ty)) => (Axclass.unoverload_const thy c_ty, ty) |
78139 | 650 |
| _ => raise BAD_THM "Equation not headed by constant"; |
651 |
fun check _ (Abs _) = raise BAD_THM "Abstraction on left hand side of equation" |
|
31156 | 652 |
| check 0 (Var _) = () |
78139 | 653 |
| check _ (Var _) = raise BAD_THM "Variable with application on left hand side of equation" |
31156 | 654 |
| check n (t1 $ t2) = (check (n+1) t1; check 0 t2) |
34894 | 655 |
| check n (Const (c_ty as (c, ty))) = |
35226 | 656 |
if allow_pats then let |
51685
385ef6706252
more standard module name Axclass (according to file name);
wenzelm
parents:
51584
diff
changeset
|
657 |
val c' = Axclass.unoverload_const thy c_ty |
45987 | 658 |
in if n = (length o binder_types) ty |
35226 | 659 |
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
|
660 |
then () |
78139 | 661 |
else raise BAD_THM (quote c ^ " is not a constructor, on left hand side of equation") |
662 |
else raise BAD_THM ("Partially applied constant " ^ quote c ^ " on left hand side of equation") |
|
663 |
end else raise BAD_THM ("Pattern not allowed here, but constant " ^ quote c ^ " encountered on left hand side of equation") |
|
31156 | 664 |
val _ = map (check 0) args; |
35226 | 665 |
val _ = if allow_nonlinear orelse is_linear thm then () |
78139 | 666 |
else raise 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
|
667 |
val _ = if (is_none o Axclass.class_of_param thy) c then () |
78139 | 668 |
else raise BAD_THM "Overloaded constant as head in equation"; |
34894 | 669 |
val _ = if not (is_constr thy c) then () |
78139 | 670 |
else raise BAD_THM "Constructor as head in equation"; |
35226 | 671 |
val _ = if not (is_abstr thy c) then () |
78139 | 672 |
else raise 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
|
673 |
val _ = check_decl_ty thy (c, ty); |
66310
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
674 |
val _ = case strip_type ty of |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
675 |
(Type (tyco, _) :: _, _) => (case lookup_vs_type_spec thy tyco of |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
676 |
SOME (_, type_spec) => (case projection_of type_spec of |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
677 |
SOME proj => |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
678 |
if c = proj |
78139 | 679 |
then raise BAD_THM "Projection as head in equation" |
66310
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
680 |
else () |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
681 |
| _ => ()) |
52475
445ae7a4e4e1
sort out code equations headed by a projection of a abstract datatype
haftmann
parents:
51717
diff
changeset
|
682 |
| _ => ()) |
445ae7a4e4e1
sort out code equations headed by a projection of a abstract datatype
haftmann
parents:
51717
diff
changeset
|
683 |
| _ => (); |
35226 | 684 |
in () end; |
685 |
||
63242
9986559617ee
clear distinction between different situations concerning strictness of code equations
haftmann
parents:
63241
diff
changeset
|
686 |
local |
9986559617ee
clear distinction between different situations concerning strictness of code equations
haftmann
parents:
63241
diff
changeset
|
687 |
|
9986559617ee
clear distinction between different situations concerning strictness of code equations
haftmann
parents:
63241
diff
changeset
|
688 |
fun raw_assert_eqn thy check_patterns (thm, proper) = |
35226 | 689 |
let |
690 |
val (lhs, rhs) = (Logic.dest_equals o Thm.plain_prop_of) thm |
|
78139 | 691 |
handle TERM _ => raise BAD_THM "Not an equation" |
692 |
| THM _ => raise BAD_THM "Not a proper equation"; |
|
66310
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
693 |
val _ = check_eqn thy {allow_nonlinear = not proper, |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
694 |
allow_consts = not (proper andalso check_patterns), allow_pats = true} thm (lhs, rhs); |
31156 | 695 |
in (thm, proper) end; |
696 |
||
63242
9986559617ee
clear distinction between different situations concerning strictness of code equations
haftmann
parents:
63241
diff
changeset
|
697 |
fun raw_assert_abs_eqn thy some_tyco thm = |
35226 | 698 |
let |
699 |
val (full_lhs, rhs) = (Logic.dest_equals o Thm.plain_prop_of) thm |
|
78139 | 700 |
handle TERM _ => raise BAD_THM "Not an equation" |
701 |
| THM _ => raise BAD_THM "Not a proper equation"; |
|
66310
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
702 |
val (proj_t, lhs) = dest_comb full_lhs |
78139 | 703 |
handle TERM _ => raise BAD_THM "Not an abstract equation"; |
66310
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
704 |
val (proj, ty) = dest_Const proj_t |
78139 | 705 |
handle TERM _ => raise 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
|
706 |
val (tyco, Ts) = (dest_Type o domain_type) ty |
78139 | 707 |
handle TERM _ => raise BAD_THM "Not an abstract equation" |
708 |
| TYPE _ => raise BAD_THM "Not an abstract equation"; |
|
35226 | 709 |
val _ = case some_tyco of SOME tyco' => if tyco = tyco' then () |
78139 | 710 |
else raise BAD_THM ("Abstract type mismatch:" ^ quote tyco ^ " vs. " ^ quote tyco') |
35226 | 711 |
| NONE => (); |
66310
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
712 |
val (vs, proj', (abs', _)) = case lookup_vs_type_spec thy tyco |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
713 |
of SOME (vs, Abstractor spec) => (vs, #projection spec, #abstractor spec) |
78139 | 714 |
| _ => raise BAD_THM ("Not an abstract type: " ^ tyco); |
66310
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
715 |
val _ = if proj = proj' then () |
78139 | 716 |
else raise BAD_THM ("Projection mismatch: " ^ quote proj ^ " vs. " ^ quote proj'); |
66310
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
717 |
val _ = check_eqn thy {allow_nonlinear = false, |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
718 |
allow_consts = false, allow_pats = false} thm (lhs, rhs); |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
719 |
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
|
720 |
else error ("Type arguments do not satisfy sort constraints of abstype certificate."); |
66310
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
721 |
in (thm, (tyco, abs')) end; |
35226 | 722 |
|
63242
9986559617ee
clear distinction between different situations concerning strictness of code equations
haftmann
parents:
63241
diff
changeset
|
723 |
in |
31156 | 724 |
|
66260 | 725 |
fun generic_assert_eqn strictness thy check_patterns eqn = |
726 |
handle_strictness fst (raw_assert_eqn thy check_patterns) strictness thy eqn; |
|
31156 | 727 |
|
63242
9986559617ee
clear distinction between different situations concerning strictness of code equations
haftmann
parents:
63241
diff
changeset
|
728 |
fun generic_assert_abs_eqn strictness thy check_patterns thm = |
9986559617ee
clear distinction between different situations concerning strictness of code equations
haftmann
parents:
63241
diff
changeset
|
729 |
handle_strictness I (raw_assert_abs_eqn thy check_patterns) strictness thy thm; |
52637
1501ebe39711
attribute "code" declares concrete and abstract code equations uniformly; added explicit "code equation" instead
haftmann
parents:
52475
diff
changeset
|
730 |
|
63242
9986559617ee
clear distinction between different situations concerning strictness of code equations
haftmann
parents:
63241
diff
changeset
|
731 |
end; |
9986559617ee
clear distinction between different situations concerning strictness of code equations
haftmann
parents:
63241
diff
changeset
|
732 |
|
9986559617ee
clear distinction between different situations concerning strictness of code equations
haftmann
parents:
63241
diff
changeset
|
733 |
fun assert_eqn thy = the o generic_assert_eqn Strict thy true; |
9986559617ee
clear distinction between different situations concerning strictness of code equations
haftmann
parents:
63241
diff
changeset
|
734 |
|
9986559617ee
clear distinction between different situations concerning strictness of code equations
haftmann
parents:
63241
diff
changeset
|
735 |
fun assert_abs_eqn thy some_tyco = the o generic_assert_abs_eqn Strict thy some_tyco; |
35226 | 736 |
|
33940
317933ce3712
crude support for type aliasses and corresponding constant signatures
haftmann
parents:
33756
diff
changeset
|
737 |
val head_eqn = dest_Const o fst o strip_comb o fst o Logic.dest_equals o Thm.plain_prop_of; |
31156 | 738 |
|
31957 | 739 |
fun const_typ_eqn thy thm = |
31156 | 740 |
let |
32640
ba6531df2c64
corrected order of type variables in code equations; more precise certificate for cases
haftmann
parents:
32544
diff
changeset
|
741 |
val (c, ty) = head_eqn thm; |
51685
385ef6706252
more standard module name Axclass (according to file name);
wenzelm
parents:
51584
diff
changeset
|
742 |
val c' = Axclass.unoverload_const thy (c, ty); |
33940
317933ce3712
crude support for type aliasses and corresponding constant signatures
haftmann
parents:
33756
diff
changeset
|
743 |
(*permissive wrt. to overloaded constants!*) |
31156 | 744 |
in (c', ty) end; |
33940
317933ce3712
crude support for type aliasses and corresponding constant signatures
haftmann
parents:
33756
diff
changeset
|
745 |
|
31957 | 746 |
fun const_eqn thy = fst o const_typ_eqn thy; |
31156 | 747 |
|
51685
385ef6706252
more standard module name Axclass (according to file name);
wenzelm
parents:
51584
diff
changeset
|
748 |
fun const_abs_eqn thy = Axclass.unoverload_const thy o dest_Const o fst o strip_comb o snd |
35226 | 749 |
o dest_comb o fst o Logic.dest_equals o Thm.plain_prop_of; |
750 |
||
751 |
fun mk_proj tyco vs ty abs rep = |
|
752 |
let |
|
753 |
val ty_abs = Type (tyco, map TFree vs); |
|
754 |
val xarg = Var (("x", 0), ty); |
|
755 |
in Logic.mk_equals (Const (rep, ty_abs --> ty) $ (Const (abs, ty --> ty_abs) $ xarg), xarg) end; |
|
756 |
||
34895 | 757 |
|
758 |
(* technical transformations of code equations *) |
|
759 |
||
63242
9986559617ee
clear distinction between different situations concerning strictness of code equations
haftmann
parents:
63241
diff
changeset
|
760 |
fun meta_rewrite thy = Local_Defs.meta_rewrite_rule (Proof_Context.init_global thy); |
9986559617ee
clear distinction between different situations concerning strictness of code equations
haftmann
parents:
63241
diff
changeset
|
761 |
|
34895 | 762 |
fun same_arity thy thms = |
31962 | 763 |
let |
77928 | 764 |
val lhs_rhss = map (Logic.dest_equals o Thm.plain_prop_of) thms; |
765 |
val k = fold (Integer.max o length o snd o strip_comb o fst) lhs_rhss 0; |
|
766 |
fun expand_eta (lhs, rhs) thm = |
|
767 |
let |
|
768 |
val l = k - length (snd (strip_comb lhs)); |
|
769 |
val (raw_vars, _) = Term.strip_abs_eta l rhs; |
|
770 |
val vars = burrow_fst (Name.variant_list (map (fst o fst) (Term.add_vars lhs []))) |
|
771 |
raw_vars; |
|
772 |
fun expand (v, ty) thm = Drule.fun_cong_rule thm |
|
773 |
(Thm.global_cterm_of thy (Var ((v, 0), ty))); |
|
774 |
in |
|
775 |
thm |
|
776 |
|> fold expand vars |
|
777 |
|> Conv.fconv_rule Drule.beta_eta_conversion |
|
778 |
end; |
|
779 |
in map2 expand_eta lhs_rhss thms end; |
|
34895 | 780 |
|
781 |
fun mk_desymbolization pre post mk vs = |
|
782 |
let |
|
783 |
val names = map (pre o fst o fst) vs |
|
56811 | 784 |
|> map (Name.desymbolize (SOME false)) |
34895 | 785 |
|> Name.variant_list [] |
786 |
|> map post; |
|
787 |
in map_filter (fn (((v, i), x), v') => |
|
788 |
if v = v' andalso i = 0 then NONE |
|
789 |
else SOME (((v, i), x), mk ((v', 0), x))) (vs ~~ names) |
|
790 |
end; |
|
791 |
||
60367 | 792 |
fun desymbolize_tvars thy thms = |
34895 | 793 |
let |
75353 | 794 |
val tvs = build (fold (Term.add_tvars o Thm.prop_of) thms); |
60805 | 795 |
val instT = |
796 |
mk_desymbolization (unprefix "'") (prefix "'") (Thm.global_ctyp_of thy o TVar) tvs; |
|
74282 | 797 |
in map (Thm.instantiate (TVars.make instT, Vars.empty)) thms end; |
34895 | 798 |
|
60367 | 799 |
fun desymbolize_vars thy thm = |
34895 | 800 |
let |
801 |
val vs = Term.add_vars (Thm.prop_of thm) []; |
|
60805 | 802 |
val inst = mk_desymbolization I I (Thm.global_cterm_of thy o Var) vs; |
74282 | 803 |
in Thm.instantiate (TVars.empty, Vars.make inst) thm end; |
34895 | 804 |
|
60367 | 805 |
fun canonize_thms thy = desymbolize_tvars thy #> same_arity thy #> map (desymbolize_vars thy); |
31156 | 806 |
|
34874 | 807 |
|
66251
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
808 |
(* preparation and classification of code equations *) |
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
809 |
|
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
810 |
fun prep_eqn strictness thy = |
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
811 |
apfst (meta_rewrite thy) |
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
812 |
#> generic_assert_eqn strictness thy false |
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
813 |
#> Option.map (fn eqn => (const_eqn thy (fst eqn), eqn)); |
67032
ed499d1252fc
strip some trailing spaces to force Pure rebuild after ce6454669360
Lars Hupel <lars.hupel@mytum.de>
parents:
66332
diff
changeset
|
814 |
|
66251
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
815 |
fun prep_eqns strictness thy = |
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
816 |
map_filter (prep_eqn strictness thy) |
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
817 |
#> AList.group (op =); |
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
818 |
|
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
819 |
fun prep_abs_eqn strictness thy = |
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
820 |
meta_rewrite thy |
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
821 |
#> generic_assert_abs_eqn strictness thy NONE |
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
822 |
#> Option.map (fn abs_eqn => (const_abs_eqn thy (fst abs_eqn), abs_eqn)); |
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
823 |
|
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
824 |
fun prep_maybe_abs_eqn thy raw_thm = |
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
825 |
let |
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
826 |
val thm = meta_rewrite thy raw_thm; |
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
827 |
val some_abs_thm = generic_assert_abs_eqn Silent thy NONE thm; |
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
828 |
in case some_abs_thm of |
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
829 |
SOME (thm, tyco) => SOME (const_abs_eqn thy thm, ((thm, true), SOME tyco)) |
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
830 |
| NONE => generic_assert_eqn Liberal thy false (thm, false) |
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
831 |
|> Option.map (fn (thm, _) => (const_eqn thy thm, ((thm, is_linear thm), NONE))) |
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
832 |
end; |
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
833 |
|
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
834 |
|
36112
7fa17a225852
user interface for abstract datatypes is an attribute, not a command
haftmann
parents:
35845
diff
changeset
|
835 |
(* abstype certificates *) |
7fa17a225852
user interface for abstract datatypes is an attribute, not a command
haftmann
parents:
35845
diff
changeset
|
836 |
|
63242
9986559617ee
clear distinction between different situations concerning strictness of code equations
haftmann
parents:
63241
diff
changeset
|
837 |
local |
9986559617ee
clear distinction between different situations concerning strictness of code equations
haftmann
parents:
63241
diff
changeset
|
838 |
|
9986559617ee
clear distinction between different situations concerning strictness of code equations
haftmann
parents:
63241
diff
changeset
|
839 |
fun raw_abstype_cert thy proto_thm = |
36112
7fa17a225852
user interface for abstract datatypes is an attribute, not a command
haftmann
parents:
35845
diff
changeset
|
840 |
let |
63073
413184c7a2a2
clarified context, notably for internal use of Simplifier;
wenzelm
parents:
61268
diff
changeset
|
841 |
val thm = (Axclass.unoverload (Proof_Context.init_global 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
|
842 |
val (lhs, rhs) = Logic.dest_equals (Thm.plain_prop_of thm) |
78139 | 843 |
handle TERM _ => raise BAD_THM "Not an equation" |
844 |
| THM _ => raise 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
|
845 |
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
|
846 |
o apfst dest_Const o dest_comb) lhs |
78139 | 847 |
handle TERM _ => raise BAD_THM "Not an abstype certificate"; |
59058
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents:
58666
diff
changeset
|
848 |
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
|
849 |
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
|
850 |
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
|
851 |
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
|
852 |
val _ = if length (binder_types raw_ty) = 1 |
04aeda922be2
explicit check for correct number of arguments for abstract constructor
haftmann
parents:
47555
diff
changeset
|
853 |
then () |
78139 | 854 |
else raise 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
|
855 |
val _ = (fst o dest_Var) param |
78139 | 856 |
handle TERM _ => raise BAD_THM "Not an abstype certificate"; |
857 |
val _ = if param = rhs then () else raise 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
|
858 |
val ((tyco, sorts), (abs, (vs, ty'))) = |
54603
89d5b69e5a08
prefer name-normalizing devarify over unvarifyT whenever appropriate
haftmann
parents:
53171
diff
changeset
|
859 |
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
|
860 |
val ty = domain_type ty'; |
49534
791e6fc53f73
more strict typscheme_equiv check: must fix variables of more specific type;
haftmann
parents:
49533
diff
changeset
|
861 |
val (vs', _) = typscheme thy (abs, ty'); |
40726
16dcfedc4eb7
keep type variable arguments of datatype constructors in bookkeeping
haftmann
parents:
40564
diff
changeset
|
862 |
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
|
863 |
|
63242
9986559617ee
clear distinction between different situations concerning strictness of code equations
haftmann
parents:
63241
diff
changeset
|
864 |
in |
9986559617ee
clear distinction between different situations concerning strictness of code equations
haftmann
parents:
63241
diff
changeset
|
865 |
|
9986559617ee
clear distinction between different situations concerning strictness of code equations
haftmann
parents:
63241
diff
changeset
|
866 |
fun check_abstype_cert strictness thy proto_thm = |
9986559617ee
clear distinction between different situations concerning strictness of code equations
haftmann
parents:
63241
diff
changeset
|
867 |
handle_strictness I (raw_abstype_cert thy) strictness thy proto_thm; |
9986559617ee
clear distinction between different situations concerning strictness of code equations
haftmann
parents:
63241
diff
changeset
|
868 |
|
9986559617ee
clear distinction between different situations concerning strictness of code equations
haftmann
parents:
63241
diff
changeset
|
869 |
end; |
9986559617ee
clear distinction between different situations concerning strictness of code equations
haftmann
parents:
63241
diff
changeset
|
870 |
|
36112
7fa17a225852
user interface for abstract datatypes is an attribute, not a command
haftmann
parents:
35845
diff
changeset
|
871 |
|
34874 | 872 |
(* code equation certificates *) |
873 |
||
34895 | 874 |
fun build_head thy (c, ty) = |
59621
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents:
59458
diff
changeset
|
875 |
Thm.global_cterm_of thy (Logic.mk_equals (Free ("HEAD", ty), Const (c, ty))); |
34874 | 876 |
|
34895 | 877 |
fun get_head thy cert_thm = |
878 |
let |
|
60949 | 879 |
val [head] = Thm.chyps_of cert_thm; |
34895 | 880 |
val (_, Const (c, ty)) = (Logic.dest_equals o Thm.term_of) head; |
881 |
in (typscheme thy (c, ty), head) end; |
|
882 |
||
35226 | 883 |
fun typscheme_projection thy = |
884 |
typscheme thy o dest_Const o fst o dest_comb o fst o Logic.dest_equals; |
|
885 |
||
886 |
fun typscheme_abs thy = |
|
887 |
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; |
|
888 |
||
889 |
fun constrain_thm thy vs sorts thm = |
|
890 |
let |
|
891 |
val mapping = map2 (fn (v, sort) => fn sort' => |
|
892 |
(v, Sorts.inter_sort (Sign.classes_of thy) (sort, sort'))) vs sorts; |
|
74282 | 893 |
val instT = |
894 |
TVars.build |
|
895 |
(fold2 (fn (v, sort) => fn (_, sort') => |
|
896 |
TVars.add (((v, 0), sort), Thm.global_ctyp_of thy (TFree (v, sort')))) vs mapping); |
|
66130 | 897 |
val subst = (Term.map_types o map_type_tfree) |
40803 | 898 |
(fn (v, _) => TFree (v, the (AList.lookup (op =) mapping v))); |
35226 | 899 |
in |
900 |
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
|
901 |
|> Thm.varifyT_global |
74282 | 902 |
|> Thm.instantiate (instT, Vars.empty) |
35226 | 903 |
|> pair subst |
904 |
end; |
|
905 |
||
906 |
fun concretify_abs thy tyco abs_thm = |
|
907 |
let |
|
66310
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
908 |
val (_, {abstractor = (c_abs, _), abs_rep, ...}) = get_abstype_spec thy tyco; |
35226 | 909 |
val lhs = (fst o Logic.dest_equals o Thm.prop_of) abs_thm |
910 |
val ty = fastype_of lhs; |
|
911 |
val ty_abs = (fastype_of o snd o dest_comb) lhs; |
|
66310
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
912 |
val abs = Thm.global_cterm_of thy (Const (c_abs, ty --> ty_abs)); |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
913 |
val raw_concrete_thm = Drule.transitive_thm OF [Thm.symmetric abs_rep, Thm.combination (Thm.reflexive abs) abs_thm]; |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
914 |
in (c_abs, (Thm.varifyT_global o zero_var_indexes) raw_concrete_thm) end; |
35226 | 915 |
|
916 |
fun add_rhss_of_eqn thy t = |
|
917 |
let |
|
45987 | 918 |
val (args, rhs) = (apfst (snd o strip_comb) o Logic.dest_equals) t; |
35226 | 919 |
fun add_const (Const (c, ty)) = insert (op =) (c, Sign.const_typargs thy (c, ty)) |
920 |
| 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
|
921 |
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
|
922 |
in add_consts rhs o fold add_consts args end; |
35226 | 923 |
|
46513
2659ee0128c2
more explicit error on malformed abstract equation; dropped dead code; tuned signature
haftmann
parents:
45987
diff
changeset
|
924 |
val dest_eqn = apfst (snd o strip_comb) o Logic.dest_equals o Logic.unvarify_global; |
35226 | 925 |
|
54889
4121d64fde90
explicit distinction between empty code equations and no code equations, including convenient declaration attributes
haftmann
parents:
54888
diff
changeset
|
926 |
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
|
927 |
| Equations of thm * bool list |
35226 | 928 |
| Projection of term * string |
929 |
| Abstract of thm * string |
|
930 |
with |
|
34891
99b9a6290446
code certificates as integral part of code generation
haftmann
parents:
34874
diff
changeset
|
931 |
|
55364 | 932 |
fun dummy_thm ctxt c = |
34891
99b9a6290446
code certificates as integral part of code generation
haftmann
parents:
34874
diff
changeset
|
933 |
let |
55364 | 934 |
val thy = Proof_Context.theory_of ctxt; |
54603
89d5b69e5a08
prefer name-normalizing devarify over unvarifyT whenever appropriate
haftmann
parents:
53171
diff
changeset
|
935 |
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
|
936 |
val (vs, _) = typscheme thy (c, raw_ty); |
51685
385ef6706252
more standard module name Axclass (according to file name);
wenzelm
parents:
51584
diff
changeset
|
937 |
val sortargs = case Axclass.class_of_param thy c |
40761
1ef64dcb24b7
corrected: use canonical variables of type scheme uniformly
haftmann
parents:
40758
diff
changeset
|
938 |
of SOME class => [[class]] |
1ef64dcb24b7
corrected: use canonical variables of type scheme uniformly
haftmann
parents:
40758
diff
changeset
|
939 |
| NONE => (case get_type_of_constr_or_abstr thy c |
1ef64dcb24b7
corrected: use canonical variables of type scheme uniformly
haftmann
parents:
40758
diff
changeset
|
940 |
of SOME (tyco, _) => (map snd o fst o the) |
1ef64dcb24b7
corrected: use canonical variables of type scheme uniformly
haftmann
parents:
40758
diff
changeset
|
941 |
(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
|
942 |
| NONE => replicate (length vs) []); |
1ef64dcb24b7
corrected: use canonical variables of type scheme uniformly
haftmann
parents:
40758
diff
changeset
|
943 |
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
|
944 |
val ty = map_type_tfree (fn (v, _) => TFree (v, the_sort v)) raw_ty |
34895 | 945 |
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
|
946 |
in Thm.weaken chead Drule.dummy_thm end; |
34891
99b9a6290446
code certificates as integral part of code generation
haftmann
parents:
34874
diff
changeset
|
947 |
|
55364 | 948 |
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
|
949 |
|
55364 | 950 |
fun cert_of_eqns ctxt c [] = Equations (dummy_thm ctxt c, []) |
66167 | 951 |
| cert_of_eqns ctxt c raw_eqns = |
34874 | 952 |
let |
55364 | 953 |
val thy = Proof_Context.theory_of ctxt; |
34895 | 954 |
val eqns = burrow_fst (canonize_thms thy) raw_eqns; |
63240 | 955 |
val _ = map (assert_eqn thy) eqns; |
34891
99b9a6290446
code certificates as integral part of code generation
haftmann
parents:
34874
diff
changeset
|
956 |
val (thms, propers) = split_list eqns; |
34895 | 957 |
val _ = map (fn thm => if c = const_eqn thy thm then () |
958 |
else error ("Wrong head of code equation,\nexpected constant " |
|
61268 | 959 |
^ string_of_const thy c ^ "\n" ^ Thm.string_of_thm_global thy thm)) thms; |
74235 | 960 |
val tvars_of = build_rev o Term.add_tvarsT; |
34891
99b9a6290446
code certificates as integral part of code generation
haftmann
parents:
34874
diff
changeset
|
961 |
val vss = map (tvars_of o snd o head_eqn) thms; |
75353 | 962 |
val inter_sorts = |
963 |
build o fold (curry (Sorts.inter_sort (Sign.classes_of thy)) o snd); |
|
34891
99b9a6290446
code certificates as integral part of code generation
haftmann
parents:
34874
diff
changeset
|
964 |
val sorts = map_transpose inter_sorts vss; |
43329
84472e198515
tuned signature: Name.invent and Name.invent_names;
wenzelm
parents:
43326
diff
changeset
|
965 |
val vts = Name.invent_names Name.context Name.aT sorts; |
74282 | 966 |
fun instantiate vs = |
967 |
Thm.instantiate (TVars.make (vs ~~ map (Thm.ctyp_of ctxt o TFree) vts), Vars.empty); |
|
968 |
val thms' = map2 instantiate vss thms; |
|
40758
4f2c3e842ef8
consider sort constraints for datatype constructors when constructing the empty equation certificate;
haftmann
parents:
40726
diff
changeset
|
969 |
val head_thm = Thm.symmetric (Thm.assume (build_head thy (head_eqn (hd thms')))); |
34874 | 970 |
fun head_conv ct = if can Thm.dest_comb ct |
971 |
then Conv.fun_conv head_conv ct |
|
972 |
else Conv.rewr_conv head_thm ct; |
|
973 |
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
|
974 |
val cert_thm = Conjunction.intr_balanced (map rewrite_head thms'); |
35226 | 975 |
in Equations (cert_thm, propers) end; |
34891
99b9a6290446
code certificates as integral part of code generation
haftmann
parents:
34874
diff
changeset
|
976 |
|
66310
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
977 |
fun cert_of_proj ctxt proj tyco = |
35226 | 978 |
let |
66127 | 979 |
val thy = Proof_Context.theory_of ctxt |
66310
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
980 |
val (vs, {abstractor = (abs, (_, ty)), projection = proj', ...}) = get_abstype_spec thy tyco; |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
981 |
val _ = if proj = proj' then () else |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
982 |
error ("Wrong head of projection,\nexpected constant " ^ string_of_const thy proj); |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
983 |
in Projection (mk_proj tyco vs ty abs proj, tyco) end; |
35226 | 984 |
|
66127 | 985 |
fun cert_of_abs ctxt tyco c raw_abs_thm = |
34874 | 986 |
let |
66127 | 987 |
val thy = Proof_Context.theory_of ctxt; |
35226 | 988 |
val abs_thm = singleton (canonize_thms thy) raw_abs_thm; |
989 |
val _ = assert_abs_eqn thy (SOME tyco) abs_thm; |
|
990 |
val _ = if c = const_abs_eqn thy abs_thm then () |
|
991 |
else error ("Wrong head of abstract code equation,\nexpected constant " |
|
61268 | 992 |
^ string_of_const thy c ^ "\n" ^ Thm.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
|
993 |
in Abstract (Thm.legacy_freezeT abs_thm, tyco) end; |
34874 | 994 |
|
54889
4121d64fde90
explicit distinction between empty code equations and no code equations, including convenient declaration attributes
haftmann
parents:
54888
diff
changeset
|
995 |
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
|
996 |
let |
4121d64fde90
explicit distinction between empty code equations and no code equations, including convenient declaration attributes
haftmann
parents:
54888
diff
changeset
|
997 |
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
|
998 |
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
|
999 |
|> Thm.implies_intr head |
4121d64fde90
explicit distinction between empty code equations and no code equations, including convenient declaration attributes
haftmann
parents:
54888
diff
changeset
|
1000 |
|> 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
|
1001 |
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
|
1002 |
|> subst |
59621
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents:
59458
diff
changeset
|
1003 |
|> 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
|
1004 |
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
|
1005 |
|> 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
|
1006 |
in cert_thm'' end; |
4121d64fde90
explicit distinction between empty code equations and no code equations, including convenient declaration attributes
haftmann
parents:
54888
diff
changeset
|
1007 |
|
4121d64fde90
explicit distinction between empty code equations and no code equations, including convenient declaration attributes
haftmann
parents:
54888
diff
changeset
|
1008 |
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
|
1009 |
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
|
1010 |
| 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
|
1011 |
Equations (constrain_cert_thm thy sorts cert_thm, propers) |
66324 | 1012 |
| constrain_cert _ _ (cert as Projection _) = |
35226 | 1013 |
cert |
1014 |
| constrain_cert thy sorts (Abstract (abs_thm, tyco)) = |
|
1015 |
Abstract (snd (constrain_thm thy (fst (typscheme_abs thy abs_thm)) sorts abs_thm), tyco); |
|
1016 |
||
54889
4121d64fde90
explicit distinction between empty code equations and no code equations, including convenient declaration attributes
haftmann
parents:
54888
diff
changeset
|
1017 |
fun conclude_cert (Nothing cert_thm) = |
78056 | 1018 |
Nothing (Thm.close_derivation \<^here> cert_thm |> Thm.trim_context) |
54889
4121d64fde90
explicit distinction between empty code equations and no code equations, including convenient declaration attributes
haftmann
parents:
54888
diff
changeset
|
1019 |
| conclude_cert (Equations (cert_thm, propers)) = |
78056 | 1020 |
Equations (Thm.close_derivation \<^here> cert_thm |> Thm.trim_context, propers) |
49971
8b50286c36d3
close code theorems explicitly after preprocessing
haftmann
parents:
49904
diff
changeset
|
1021 |
| conclude_cert (cert as Projection _) = |
8b50286c36d3
close code theorems explicitly after preprocessing
haftmann
parents:
49904
diff
changeset
|
1022 |
cert |
8b50286c36d3
close code theorems explicitly after preprocessing
haftmann
parents:
49904
diff
changeset
|
1023 |
| conclude_cert (Abstract (abs_thm, tyco)) = |
78056 | 1024 |
Abstract (Thm.close_derivation \<^here> abs_thm |> Thm.trim_context, tyco); |
49971
8b50286c36d3
close code theorems explicitly after preprocessing
haftmann
parents:
49904
diff
changeset
|
1025 |
|
54889
4121d64fde90
explicit distinction between empty code equations and no code equations, including convenient declaration attributes
haftmann
parents:
54888
diff
changeset
|
1026 |
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
|
1027 |
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
|
1028 |
| typscheme_of_cert thy (Equations (cert_thm, _)) = |
35226 | 1029 |
fst (get_head thy cert_thm) |
1030 |
| typscheme_of_cert thy (Projection (proj, _)) = |
|
1031 |
typscheme_projection thy proj |
|
1032 |
| typscheme_of_cert thy (Abstract (abs_thm, _)) = |
|
1033 |
typscheme_abs thy abs_thm; |
|
34874 | 1034 |
|
54889
4121d64fde90
explicit distinction between empty code equations and no code equations, including convenient declaration attributes
haftmann
parents:
54888
diff
changeset
|
1035 |
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
|
1036 |
let |
4121d64fde90
explicit distinction between empty code equations and no code equations, including convenient declaration attributes
haftmann
parents:
54888
diff
changeset
|
1037 |
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
|
1038 |
in (vs, []) end |
4121d64fde90
explicit distinction between empty code equations and no code equations, including convenient declaration attributes
haftmann
parents:
54888
diff
changeset
|
1039 |
| typargs_deps_of_cert thy (Equations (cert_thm, propers)) = |
35226 | 1040 |
let |
1041 |
val vs = (fst o fst) (get_head thy cert_thm); |
|
1042 |
val equations = if null propers then [] else |
|
1043 |
Thm.prop_of cert_thm |
|
1044 |
|> Logic.dest_conjunction_balanced (length propers); |
|
75353 | 1045 |
in (vs, build (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
|
1046 |
| typargs_deps_of_cert thy (Projection (t, _)) = |
35226 | 1047 |
(fst (typscheme_projection thy t), add_rhss_of_eqn thy t []) |
1048 |
| typargs_deps_of_cert thy (Abstract (abs_thm, tyco)) = |
|
1049 |
let |
|
1050 |
val vs = fst (typscheme_abs thy abs_thm); |
|
1051 |
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
|
1052 |
in (vs, add_rhss_of_eqn thy (Logic.unvarify_types_global (Thm.prop_of concrete_thm)) []) end; |
34895 | 1053 |
|
54889
4121d64fde90
explicit distinction between empty code equations and no code equations, including convenient declaration attributes
haftmann
parents:
54888
diff
changeset
|
1054 |
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
|
1055 |
(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
|
1056 |
| equations_of_cert thy (cert as Equations (cert_thm, propers)) = |
35226 | 1057 |
let |
1058 |
val tyscm = typscheme_of_cert thy cert; |
|
1059 |
val thms = if null propers then [] else |
|
1060 |
cert_thm |
|
78056 | 1061 |
|> Thm.transfer thy |
35624 | 1062 |
|> 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
|
1063 |
|> Thm.varifyT_global |
35226 | 1064 |
|> Conjunction.elim_balanced (length propers); |
77702 | 1065 |
fun abstractions (args, rhs) = (map (pair NONE) args, (NONE, rhs)); |
54889
4121d64fde90
explicit distinction between empty code equations and no code equations, including convenient declaration attributes
haftmann
parents:
54888
diff
changeset
|
1066 |
in (tyscm, SOME (map (abstractions o dest_eqn o Thm.prop_of) thms ~~ (map SOME thms ~~ propers))) end |
35226 | 1067 |
| equations_of_cert thy (Projection (t, tyco)) = |
1068 |
let |
|
66310
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1069 |
val (_, {abstractor = (abs, _), ...}) = get_abstype_spec thy tyco; |
35226 | 1070 |
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
|
1071 |
val t' = Logic.varify_types_global t; |
77702 | 1072 |
fun abstractions (args, rhs) = (map (pair (SOME abs)) args, (NONE, rhs)); |
54889
4121d64fde90
explicit distinction between empty code equations and no code equations, including convenient declaration attributes
haftmann
parents:
54888
diff
changeset
|
1073 |
in (tyscm, SOME [((abstractions o dest_eqn) t', (NONE, true))]) end |
35226 | 1074 |
| equations_of_cert thy (Abstract (abs_thm, tyco)) = |
1075 |
let |
|
1076 |
val tyscm = typscheme_abs thy abs_thm; |
|
78056 | 1077 |
val (abs, concrete_thm) = concretify_abs thy tyco (Thm.transfer thy abs_thm); |
77702 | 1078 |
fun abstractions (args, rhs) = (map (pair NONE) args, (SOME abs, rhs)); |
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
|
1079 |
in |
54889
4121d64fde90
explicit distinction between empty code equations and no code equations, including convenient declaration attributes
haftmann
parents:
54888
diff
changeset
|
1080 |
(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
|
1081 |
(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
|
1082 |
end; |
34895 | 1083 |
|
66324 | 1084 |
fun pretty_cert _ (Nothing _) = |
66332 | 1085 |
[] |
54889
4121d64fde90
explicit distinction between empty code equations and no code equations, including convenient declaration attributes
haftmann
parents:
54888
diff
changeset
|
1086 |
| pretty_cert thy (cert as Equations _) = |
63073
413184c7a2a2
clarified context, notably for internal use of Simplifier;
wenzelm
parents:
61268
diff
changeset
|
1087 |
(map_filter |
413184c7a2a2
clarified context, notably for internal use of Simplifier;
wenzelm
parents:
61268
diff
changeset
|
1088 |
(Option.map (Thm.pretty_thm_global thy o |
413184c7a2a2
clarified context, notably for internal use of Simplifier;
wenzelm
parents:
61268
diff
changeset
|
1089 |
Axclass.overload (Proof_Context.init_global 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
|
1090 |
o these o snd o equations_of_cert thy) cert |
35226 | 1091 |
| 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
|
1092 |
[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
|
1093 |
| pretty_cert thy (Abstract (abs_thm, _)) = |
63073
413184c7a2a2
clarified context, notably for internal use of Simplifier;
wenzelm
parents:
61268
diff
changeset
|
1094 |
[(Thm.pretty_thm_global thy o |
413184c7a2a2
clarified context, notably for internal use of Simplifier;
wenzelm
parents:
61268
diff
changeset
|
1095 |
Axclass.overload (Proof_Context.init_global thy) o Thm.varifyT_global) abs_thm]; |
35226 | 1096 |
|
34895 | 1097 |
end; |
34891
99b9a6290446
code certificates as integral part of code generation
haftmann
parents:
34874
diff
changeset
|
1098 |
|
34874 | 1099 |
|
57430
020cea57eaa4
tracing facilities for the code generator preprocessor
haftmann
parents:
57429
diff
changeset
|
1100 |
(* code certificate access with preprocessing *) |
35226 | 1101 |
|
48075
ec5e62b868eb
apply preprocessing simpset also to rhs of abstract code equations
haftmann
parents:
48068
diff
changeset
|
1102 |
fun eqn_conv conv ct = |
ec5e62b868eb
apply preprocessing simpset also to rhs of abstract code equations
haftmann
parents:
48068
diff
changeset
|
1103 |
let |
ec5e62b868eb
apply preprocessing simpset also to rhs of abstract code equations
haftmann
parents:
48068
diff
changeset
|
1104 |
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
|
1105 |
then Conv.combination_conv lhs_conv conv ct |
ec5e62b868eb
apply preprocessing simpset also to rhs of abstract code equations
haftmann
parents:
48068
diff
changeset
|
1106 |
else Conv.all_conv ct; |
ec5e62b868eb
apply preprocessing simpset also to rhs of abstract code equations
haftmann
parents:
48068
diff
changeset
|
1107 |
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
|
1108 |
|
55364 | 1109 |
fun rewrite_eqn conv ctxt = |
1110 |
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
|
1111 |
|
74978 | 1112 |
fun apply_functrans ctxt functrans = |
58559
d230e7075bcf
code preprocessor tracing also for function transformers
haftmann
parents:
58011
diff
changeset
|
1113 |
let |
d230e7075bcf
code preprocessor tracing also for function transformers
haftmann
parents:
58011
diff
changeset
|
1114 |
fun trace_eqns s eqns = (Pretty.writeln o Pretty.chunks) |
61268 | 1115 |
(Pretty.str s :: map (Thm.pretty_thm ctxt o fst) eqns); |
58559
d230e7075bcf
code preprocessor tracing also for function transformers
haftmann
parents:
58011
diff
changeset
|
1116 |
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
|
1117 |
in |
d230e7075bcf
code preprocessor tracing also for function transformers
haftmann
parents:
58011
diff
changeset
|
1118 |
tap (tracing "before function transformation") |
d230e7075bcf
code preprocessor tracing also for function transformers
haftmann
parents:
58011
diff
changeset
|
1119 |
#> (perhaps o perhaps_loop o perhaps_apply) functrans |
d230e7075bcf
code preprocessor tracing also for function transformers
haftmann
parents:
58011
diff
changeset
|
1120 |
#> tap (tracing "after function transformation") |
d230e7075bcf
code preprocessor tracing also for function transformers
haftmann
parents:
58011
diff
changeset
|
1121 |
end; |
48075
ec5e62b868eb
apply preprocessing simpset also to rhs of abstract code equations
haftmann
parents:
48068
diff
changeset
|
1122 |
|
74978 | 1123 |
fun preprocess conv ctxt = |
1124 |
rewrite_eqn conv ctxt |
|
1125 |
#> Axclass.unoverload ctxt; |
|
1126 |
||
57429 | 1127 |
fun get_cert ctxt functrans c = |
66310
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1128 |
case lookup_proper_fun_spec (specs_of (Proof_Context.theory_of ctxt)) c of |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1129 |
NONE => nothing_cert ctxt c |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1130 |
| SOME (Eqns (_, eqns)) => eqns |
74978 | 1131 |
|> (map o apfst) (Thm.transfer' ctxt) |
1132 |
|> apply_functrans ctxt functrans |
|
1133 |
|> (map o apfst) (preprocess eqn_conv ctxt) |
|
1134 |
|> cert_of_eqns ctxt c |
|
66310
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1135 |
| SOME (Proj (_, (tyco, _))) => cert_of_proj ctxt c tyco |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1136 |
| SOME (Abstr (abs_thm, (tyco, _))) => abs_thm |
74978 | 1137 |
|> Thm.transfer' ctxt |
1138 |
|> preprocess Conv.arg_conv ctxt |
|
1139 |
|> cert_of_abs ctxt tyco c; |
|
31962 | 1140 |
|
1141 |
||
66128 | 1142 |
(* case certificates *) |
31156 | 1143 |
|
66128 | 1144 |
local |
1145 |
||
1146 |
fun raw_case_cert thm = |
|
31156 | 1147 |
let |
1148 |
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
|
1149 |
o apsnd Logic.dest_conjunctions o Logic.dest_implies o Thm.plain_prop_of) thm; |
66132 | 1150 |
val _ = case head of Free _ => () |
1151 |
| Var _ => () |
|
31156 | 1152 |
| _ => raise TERM ("case_cert", []); |
1153 |
val ([(case_var, _)], case_expr) = Term.strip_abs_eta 1 raw_case_expr; |
|
1154 |
val (Const (case_const, _), raw_params) = strip_comb case_expr; |
|
1155 |
val n = find_index (fn Free (v, _) => v = case_var | _ => false) raw_params; |
|
1156 |
val _ = if n = ~1 then raise TERM ("case_cert", []) else (); |
|
1157 |
val params = map (fst o dest_Var) (nth_drop n raw_params); |
|
1158 |
fun dest_case t = |
|
1159 |
let |
|
1160 |
val (head' $ t_co, rhs) = Logic.dest_equals t; |
|
1161 |
val _ = if head' = head then () else raise TERM ("case_cert", []); |
|
1162 |
val (Const (co, _), args) = strip_comb t_co; |
|
1163 |
val (Var (param, _), args') = strip_comb rhs; |
|
1164 |
val _ = if args' = args then () else raise TERM ("case_cert", []); |
|
1165 |
in (param, co) end; |
|
1166 |
fun analyze_cases cases = |
|
1167 |
let |
|
75353 | 1168 |
val co_list = build (fold (AList.update (op =) o dest_case) cases); |
47437
4625ee486ff6
generalise case certificates to allow ignored parameters
Andreas Lochbihler
parents:
46513
diff
changeset
|
1169 |
in map (AList.lookup (op =) co_list) params end; |
31156 | 1170 |
fun analyze_let t = |
1171 |
let |
|
1172 |
val (head' $ arg, Var (param', _) $ arg') = Logic.dest_equals t; |
|
1173 |
val _ = if head' = head then () else raise TERM ("case_cert", []); |
|
1174 |
val _ = if arg' = arg then () else raise TERM ("case_cert", []); |
|
1175 |
val _ = if [param'] = params then () else raise TERM ("case_cert", []); |
|
1176 |
in [] end; |
|
1177 |
fun analyze (cases as [let_case]) = |
|
1178 |
(analyze_cases cases handle Bind => analyze_let let_case) |
|
1179 |
| analyze cases = analyze_cases cases; |
|
1180 |
in (case_const, (n, analyze cases)) end; |
|
1181 |
||
66128 | 1182 |
in |
1183 |
||
1184 |
fun case_cert thm = raw_case_cert thm |
|
31156 | 1185 |
handle Bind => error "bad case certificate" |
1186 |
| TERM _ => error "bad case certificate"; |
|
1187 |
||
66128 | 1188 |
end; |
1189 |
||
66310
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1190 |
fun lookup_case_spec thy = History.lookup ((cases_of o specs_of) thy); |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1191 |
|
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1192 |
fun get_case_schema thy c = case lookup_case_spec thy c of |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1193 |
SOME (Case {schema, ...}) => SOME schema |
66189
23917e861eaa
treat "undefined" constants internally as special form of case combinators
haftmann
parents:
66167
diff
changeset
|
1194 |
| _ => NONE; |
24219 | 1195 |
|
66310
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1196 |
fun get_case_cong thy c = case lookup_case_spec thy c of |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1197 |
SOME (Case {cong, ...}) => SOME cong |
66189
23917e861eaa
treat "undefined" constants internally as special form of case combinators
haftmann
parents:
66167
diff
changeset
|
1198 |
| _ => NONE; |
23917e861eaa
treat "undefined" constants internally as special form of case combinators
haftmann
parents:
66167
diff
changeset
|
1199 |
|
66310
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1200 |
fun is_undefined thy c = case lookup_case_spec thy c of |
66189
23917e861eaa
treat "undefined" constants internally as special form of case combinators
haftmann
parents:
66167
diff
changeset
|
1201 |
SOME Undefined => true |
23917e861eaa
treat "undefined" constants internally as special form of case combinators
haftmann
parents:
66167
diff
changeset
|
1202 |
| _ => false; |
24219 | 1203 |
|
1204 |
||
31962 | 1205 |
(* diagnostic *) |
24219 | 1206 |
|
1207 |
fun print_codesetup thy = |
|
1208 |
let |
|
42360 | 1209 |
val ctxt = Proof_Context.init_global thy; |
66310
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1210 |
val specs = specs_of thy; |
35226 | 1211 |
fun pretty_equations const thms = |
51584 | 1212 |
(Pretty.block o Pretty.fbreaks) |
61268 | 1213 |
(Pretty.str (string_of_const thy const) :: map (Thm.pretty_thm_item ctxt) thms); |
66310
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1214 |
fun pretty_function (const, Eqns (_, eqns)) = |
66131 | 1215 |
pretty_equations const (map fst eqns) |
66251
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
1216 |
| pretty_function (const, Proj (proj, _)) = Pretty.block |
35226 | 1217 |
[Pretty.str (string_of_const thy const), Pretty.fbrk, Syntax.pretty_term ctxt proj] |
66251
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
1218 |
| pretty_function (const, Abstr (thm, _)) = pretty_equations const [thm]; |
35226 | 1219 |
fun pretty_typ (tyco, vs) = Pretty.str |
1220 |
(string_of_typ thy (Type (tyco, map TFree vs))); |
|
66310
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1221 |
fun pretty_type_spec (typ, (cos, abstract)) = if null cos |
35226 | 1222 |
then pretty_typ typ |
1223 |
else (Pretty.block o Pretty.breaks) ( |
|
1224 |
pretty_typ typ |
|
1225 |
:: Pretty.str "=" |
|
1226 |
:: (if abstract then [Pretty.str "(abstract)"] else []) |
|
40726
16dcfedc4eb7
keep type variable arguments of datatype constructors in bookkeeping
haftmann
parents:
40564
diff
changeset
|
1227 |
@ 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
|
1228 |
| (c, (_, tys)) => |
35226 | 1229 |
(Pretty.block o Pretty.breaks) |
1230 |
(Pretty.str (string_of_const thy c) |
|
1231 |
:: Pretty.str "of" |
|
1232 |
:: map (Pretty.quote o Syntax.pretty_typ_global thy) tys)) cos) |
|
1233 |
); |
|
66310
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1234 |
fun pretty_case_param NONE = "<ignored>" |
75399
cdf84288d93c
pass constructor arity as part of case certficiate
haftmann
parents:
75353
diff
changeset
|
1235 |
| pretty_case_param (SOME (c, _)) = string_of_const thy c |
66310
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1236 |
fun pretty_case (const, Case {schema = (_, (_, [])), ...}) = |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1237 |
Pretty.str (string_of_const thy const) |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1238 |
| pretty_case (const, Case {schema = (_, (_, cos)), ...}) = |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1239 |
(Pretty.block o Pretty.breaks) [ |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1240 |
Pretty.str (string_of_const thy const), Pretty.str "with", |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1241 |
(Pretty.block o Pretty.commas o map (Pretty.str o pretty_case_param)) cos] |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1242 |
| pretty_case (const, Undefined) = |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1243 |
(Pretty.block o Pretty.breaks) [ |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1244 |
Pretty.str (string_of_const thy const), Pretty.str "<undefined>"]; |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1245 |
val functions = all_fun_specs specs |
59058
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents:
58666
diff
changeset
|
1246 |
|> sort (string_ord o apply2 fst); |
66310
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1247 |
val types = History.all (types_of specs) |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1248 |
|> map (fn (tyco, (vs, spec)) => |
35226 | 1249 |
((tyco, vs), constructors_of spec)) |
59058
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents:
58666
diff
changeset
|
1250 |
|> sort (string_ord o apply2 (fst o fst)); |
66310
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1251 |
val cases = History.all (cases_of specs) |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1252 |
|> filter (fn (_, No_Case) => false | _ => true) |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1253 |
|> sort (string_ord o apply2 fst); |
24219 | 1254 |
in |
56334
6b3739fee456
some shortcuts for chunks, which sometimes avoid bulky string output;
wenzelm
parents:
56204
diff
changeset
|
1255 |
Pretty.writeln_chunks [ |
24219 | 1256 |
Pretty.block ( |
66310
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1257 |
Pretty.str "types:" :: Pretty.fbrk |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1258 |
:: (Pretty.fbreaks o map pretty_type_spec) types |
24219 | 1259 |
), |
25968 | 1260 |
Pretty.block ( |
66310
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1261 |
Pretty.str "functions:" :: Pretty.fbrk |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1262 |
:: (Pretty.fbreaks o map pretty_function) functions |
34901 | 1263 |
), |
1264 |
Pretty.block ( |
|
1265 |
Pretty.str "cases:" :: Pretty.fbrk |
|
1266 |
:: (Pretty.fbreaks o map pretty_case) cases |
|
24219 | 1267 |
) |
1268 |
] |
|
1269 |
end; |
|
1270 |
||
1271 |
||
66310
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1272 |
(** declaration of executable ingredients **) |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1273 |
|
66330
dcb3e6bdc00a
one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents:
66329
diff
changeset
|
1274 |
(* plugins for dependent applications *) |
dcb3e6bdc00a
one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents:
66329
diff
changeset
|
1275 |
|
dcb3e6bdc00a
one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents:
66329
diff
changeset
|
1276 |
structure Codetype_Plugin = Plugin(type T = string); |
dcb3e6bdc00a
one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents:
66329
diff
changeset
|
1277 |
|
67147 | 1278 |
val codetype_plugin = Plugin_Name.declare_setup \<^binding>\<open>codetype\<close>; |
66330
dcb3e6bdc00a
one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents:
66329
diff
changeset
|
1279 |
|
dcb3e6bdc00a
one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents:
66329
diff
changeset
|
1280 |
fun type_interpretation f = |
dcb3e6bdc00a
one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents:
66329
diff
changeset
|
1281 |
Codetype_Plugin.interpretation codetype_plugin |
dcb3e6bdc00a
one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents:
66329
diff
changeset
|
1282 |
(fn tyco => Local_Theory.background_theory |
dcb3e6bdc00a
one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents:
66329
diff
changeset
|
1283 |
(fn thy => |
dcb3e6bdc00a
one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents:
66329
diff
changeset
|
1284 |
thy |
dcb3e6bdc00a
one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents:
66329
diff
changeset
|
1285 |
|> Sign.root_path |
dcb3e6bdc00a
one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents:
66329
diff
changeset
|
1286 |
|> Sign.add_path (Long_Name.qualifier tyco) |
67032
ed499d1252fc
strip some trailing spaces to force Pure rebuild after ce6454669360
Lars Hupel <lars.hupel@mytum.de>
parents:
66332
diff
changeset
|
1287 |
|> f tyco |
66330
dcb3e6bdc00a
one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents:
66329
diff
changeset
|
1288 |
|> Sign.restore_naming thy)); |
dcb3e6bdc00a
one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents:
66329
diff
changeset
|
1289 |
|
dcb3e6bdc00a
one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents:
66329
diff
changeset
|
1290 |
fun datatype_interpretation f = |
dcb3e6bdc00a
one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents:
66329
diff
changeset
|
1291 |
type_interpretation (fn tyco => fn thy => |
dcb3e6bdc00a
one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents:
66329
diff
changeset
|
1292 |
case get_type thy tyco of |
dcb3e6bdc00a
one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents:
66329
diff
changeset
|
1293 |
(spec, false) => f (tyco, spec) thy |
dcb3e6bdc00a
one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents:
66329
diff
changeset
|
1294 |
| (_, true) => thy |
dcb3e6bdc00a
one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents:
66329
diff
changeset
|
1295 |
); |
dcb3e6bdc00a
one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents:
66329
diff
changeset
|
1296 |
|
dcb3e6bdc00a
one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents:
66329
diff
changeset
|
1297 |
fun abstype_interpretation f = |
dcb3e6bdc00a
one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents:
66329
diff
changeset
|
1298 |
type_interpretation (fn tyco => fn thy => |
dcb3e6bdc00a
one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents:
66329
diff
changeset
|
1299 |
case try (get_abstype_spec thy) tyco of |
dcb3e6bdc00a
one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents:
66329
diff
changeset
|
1300 |
SOME spec => f (tyco, spec) thy |
dcb3e6bdc00a
one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents:
66329
diff
changeset
|
1301 |
| NONE => thy |
dcb3e6bdc00a
one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents:
66329
diff
changeset
|
1302 |
); |
dcb3e6bdc00a
one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents:
66329
diff
changeset
|
1303 |
|
dcb3e6bdc00a
one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents:
66329
diff
changeset
|
1304 |
fun register_tyco_for_plugin tyco = |
dcb3e6bdc00a
one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents:
66329
diff
changeset
|
1305 |
Named_Target.theory_map (Codetype_Plugin.data_default tyco); |
dcb3e6bdc00a
one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents:
66329
diff
changeset
|
1306 |
|
dcb3e6bdc00a
one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents:
66329
diff
changeset
|
1307 |
|
66310
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1308 |
(* abstract code declarations *) |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1309 |
|
78140 | 1310 |
fun code_declaration (strictness: strictness) transform f x = |
1311 |
let val x0 = transform Morphism.trim_context_morphism x in |
|
1312 |
Local_Theory.declaration {syntax = false, pervasive = false, pos = \<^here>} |
|
1313 |
(fn phi => Context.mapping (f strictness (transform phi x0)) I) |
|
1314 |
end; |
|
66310
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1315 |
|
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1316 |
|
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1317 |
(* types *) |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1318 |
|
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1319 |
fun invalidate_constructors_of (_, type_spec) = |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1320 |
fold (fn (c, _) => History.register c unimplemented) (fst (constructors_of type_spec)); |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1321 |
|
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1322 |
fun invalidate_abstract_functions_of (_, type_spec) = |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1323 |
fold (fn c => History.register c unimplemented) (abstract_functions_of type_spec); |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1324 |
|
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1325 |
fun invalidate_case_combinators_of (_, type_spec) = |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1326 |
fold (fn c => History.register c No_Case) (case_combinators_of type_spec); |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1327 |
|
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1328 |
fun register_type (tyco, vs_typ_spec) specs = |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1329 |
let |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1330 |
val olds = the_list (History.lookup (types_of specs) tyco); |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1331 |
in |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1332 |
specs |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1333 |
|> map_functions (fold invalidate_abstract_functions_of olds |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1334 |
#> invalidate_constructors_of vs_typ_spec) |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1335 |
|> map_cases (fold invalidate_case_combinators_of olds) |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1336 |
|> map_types (History.register tyco vs_typ_spec) |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1337 |
end; |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1338 |
|
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1339 |
fun declare_datatype_global proto_constrs thy = |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1340 |
let |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1341 |
fun unoverload_const_typ (c, ty) = |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1342 |
(Axclass.unoverload_const thy (c, ty), ty); |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1343 |
val constrs = map unoverload_const_typ proto_constrs; |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1344 |
val (tyco, (vs, cos)) = constrset_of_consts thy constrs; |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1345 |
in |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1346 |
thy |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1347 |
|> modify_specs (register_type |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1348 |
(tyco, (vs, Constructors {constructors = cos, case_combinators = []}))) |
66330
dcb3e6bdc00a
one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents:
66329
diff
changeset
|
1349 |
|> register_tyco_for_plugin tyco |
66310
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1350 |
end; |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1351 |
|
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1352 |
fun declare_datatype_cmd raw_constrs thy = |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1353 |
declare_datatype_global (map (read_bare_const thy) raw_constrs) thy; |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1354 |
|
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1355 |
fun generic_declare_abstype strictness proto_thm thy = |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1356 |
case check_abstype_cert strictness thy proto_thm of |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1357 |
SOME (tyco, (vs, (abstractor as (abs, (_, ty)), (proj, abs_rep)))) => |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1358 |
thy |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1359 |
|> modify_specs (register_type |
78056 | 1360 |
(tyco, (vs, Abstractor {abstractor = abstractor, projection = proj, |
1361 |
abs_rep = Thm.trim_context abs_rep, more_abstract_functions = []})) |
|
66310
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1362 |
#> register_fun_spec proj |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1363 |
(Proj (Logic.varify_types_global (mk_proj tyco vs ty abs proj), (tyco, abs)))) |
66330
dcb3e6bdc00a
one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents:
66329
diff
changeset
|
1364 |
|> register_tyco_for_plugin tyco |
66310
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1365 |
| NONE => thy; |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1366 |
|
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1367 |
val declare_abstype_global = generic_declare_abstype Strict; |
78068 | 1368 |
val declare_abstype = code_declaration Liberal Morphism.thm generic_declare_abstype; |
66310
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1369 |
|
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1370 |
|
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1371 |
(* functions *) |
31962 | 1372 |
|
63239
d562c9948dee
explicit tagging of code equations de-baroquifies interface
haftmann
parents:
63238
diff
changeset
|
1373 |
(* |
63242
9986559617ee
clear distinction between different situations concerning strictness of code equations
haftmann
parents:
63241
diff
changeset
|
1374 |
strictness wrt. shape of theorem propositions: |
66251
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
1375 |
* default equations: silent |
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
1376 |
* using declarations and attributes: warnings (after morphism application!) |
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
1377 |
* using global declarations (... -> thy -> thy): strict |
63242
9986559617ee
clear distinction between different situations concerning strictness of code equations
haftmann
parents:
63241
diff
changeset
|
1378 |
* internal processing after storage: strict |
63239
d562c9948dee
explicit tagging of code equations de-baroquifies interface
haftmann
parents:
63238
diff
changeset
|
1379 |
*) |
d562c9948dee
explicit tagging of code equations de-baroquifies interface
haftmann
parents:
63238
diff
changeset
|
1380 |
|
66251
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
1381 |
local |
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
1382 |
|
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
1383 |
fun subsumptive_add thy verbose (thm, proper) eqns = |
52637
1501ebe39711
attribute "code" declares concrete and abstract code equations uniformly; added explicit "code equation" instead
haftmann
parents:
52475
diff
changeset
|
1384 |
let |
67522 | 1385 |
val args_of = drop_prefix is_Var o rev o snd o strip_comb |
67646 | 1386 |
o Term.map_types Type.strip_sorts o fst o Logic.dest_equals o Thm.plain_prop_of |
1387 |
o Thm.transfer thy; |
|
66251
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
1388 |
val args = args_of thm; |
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
1389 |
val incr_idx = Logic.incr_indexes ([], [], Thm.maxidx_of thm + 1); |
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
1390 |
fun matches_args args' = |
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
1391 |
let |
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
1392 |
val k = length args' - length args |
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
1393 |
in if k >= 0 |
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
1394 |
then Pattern.matchess thy (args, (map incr_idx o drop k) args') |
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
1395 |
else false |
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
1396 |
end; |
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
1397 |
fun drop (thm', proper') = if (proper orelse not proper') |
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
1398 |
andalso matches_args (args_of thm') then |
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
1399 |
(if verbose then warning ("Code generator: dropping subsumed code equation\n" ^ |
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
1400 |
Thm.string_of_thm_global thy thm') else (); true) |
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
1401 |
else false; |
70494 | 1402 |
in (thm |> Thm.close_derivation \<^here> |> Thm.trim_context, proper) :: filter_out drop eqns end; |
31962 | 1403 |
|
66310
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1404 |
fun add_eqn_for (c, eqn) thy = |
67646 | 1405 |
thy |> modify_specs (modify_pending_eqns c (subsumptive_add thy true eqn)); |
66251
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
1406 |
|
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
1407 |
fun add_eqns_for default (c, proto_eqns) thy = |
66310
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1408 |
thy |> modify_specs (fn specs => |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1409 |
if is_default (lookup_fun_spec specs c) orelse not default |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1410 |
then |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1411 |
let |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1412 |
val eqns = [] |
67646 | 1413 |
|> fold_rev (subsumptive_add thy (not default)) proto_eqns; |
66310
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1414 |
in specs |> register_fun_spec c (Eqns (default, eqns)) end |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1415 |
else specs); |
63242
9986559617ee
clear distinction between different situations concerning strictness of code equations
haftmann
parents:
63241
diff
changeset
|
1416 |
|
66310
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1417 |
fun add_abstract_for (c, (thm, tyco_abs as (tyco, _))) = |
78056 | 1418 |
modify_specs (register_fun_spec c (Abstr (Thm.close_derivation \<^here> thm |> Thm.trim_context, tyco_abs)) |
66310
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1419 |
#> map_types (History.modify_entry tyco (add_abstract_function c))) |
66251
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
1420 |
|
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
1421 |
in |
63242
9986559617ee
clear distinction between different situations concerning strictness of code equations
haftmann
parents:
63241
diff
changeset
|
1422 |
|
66251
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
1423 |
fun generic_declare_eqns default strictness raw_eqns thy = |
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
1424 |
fold (add_eqns_for default) (prep_eqns strictness thy raw_eqns) thy; |
37425 | 1425 |
|
66251
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
1426 |
fun generic_add_eqn strictness raw_eqn thy = |
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
1427 |
fold add_eqn_for (the_list (prep_eqn strictness thy raw_eqn)) thy; |
52637
1501ebe39711
attribute "code" declares concrete and abstract code equations uniformly; added explicit "code equation" instead
haftmann
parents:
52475
diff
changeset
|
1428 |
|
66251
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
1429 |
fun generic_declare_abstract_eqn strictness raw_abs_eqn thy = |
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
1430 |
fold add_abstract_for (the_list (prep_abs_eqn strictness thy raw_abs_eqn)) thy; |
55720 | 1431 |
|
63242
9986559617ee
clear distinction between different situations concerning strictness of code equations
haftmann
parents:
63241
diff
changeset
|
1432 |
fun add_maybe_abs_eqn_liberal thm thy = |
66251
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
1433 |
case prep_maybe_abs_eqn thy thm |
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
1434 |
of SOME (c, (eqn, NONE)) => add_eqn_for (c, eqn) thy |
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
1435 |
| SOME (c, ((thm, _), SOME tyco)) => add_abstract_for (c, (thm, tyco)) thy |
52637
1501ebe39711
attribute "code" declares concrete and abstract code equations uniformly; added explicit "code equation" instead
haftmann
parents:
52475
diff
changeset
|
1436 |
| NONE => thy; |
35226 | 1437 |
|
66251
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
1438 |
end; |
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
1439 |
|
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
1440 |
val declare_default_eqns_global = generic_declare_eqns true Silent; |
78068 | 1441 |
val declare_default_eqns = code_declaration Silent (map o apfst o Morphism.thm) (generic_declare_eqns true); |
66251
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
1442 |
|
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
1443 |
val declare_eqns_global = generic_declare_eqns false Strict; |
78068 | 1444 |
val declare_eqns = code_declaration Liberal (map o apfst o Morphism.thm) (generic_declare_eqns false); |
66251
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
1445 |
|
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
1446 |
val add_eqn_global = generic_add_eqn Strict; |
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
1447 |
|
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
1448 |
fun del_eqn_global thm thy = |
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
1449 |
case prep_eqn Liberal thy (thm, false) of |
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
1450 |
SOME (c, (thm, _)) => |
66310
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1451 |
modify_specs (modify_pending_eqns c (filter_out (fn (thm', _) => Thm.eq_thm_prop (thm, thm')))) thy |
31962 | 1452 |
| NONE => thy; |
1453 |
||
66251
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
1454 |
val declare_abstract_eqn_global = generic_declare_abstract_eqn Strict; |
78068 | 1455 |
val declare_abstract_eqn = code_declaration Liberal Morphism.thm generic_declare_abstract_eqn; |
54889
4121d64fde90
explicit distinction between empty code equations and no code equations, including convenient declaration attributes
haftmann
parents:
54888
diff
changeset
|
1456 |
|
67032
ed499d1252fc
strip some trailing spaces to force Pure rebuild after ce6454669360
Lars Hupel <lars.hupel@mytum.de>
parents:
66332
diff
changeset
|
1457 |
fun declare_aborting_global c = |
66310
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1458 |
modify_specs (register_fun_spec c aborting); |
66251
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
1459 |
|
66310
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1460 |
fun declare_unimplemented_global c = |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1461 |
modify_specs (register_fun_spec c unimplemented); |
34244 | 1462 |
|
1463 |
||
1464 |
(* cases *) |
|
1465 |
||
40758
4f2c3e842ef8
consider sort constraints for datatype constructors when constructing the empty equation certificate;
haftmann
parents:
40726
diff
changeset
|
1466 |
fun case_cong thy case_const (num_args, (pos, _)) = |
37438 | 1467 |
let |
43326
47cf4bc789aa
simplified Name.variant -- discontinued builtin fold_map;
wenzelm
parents:
42707
diff
changeset
|
1468 |
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
|
1469 |
val (zs, _) = fold_map Name.variant (replicate (num_args - 1) "") ctxt; |
37438 | 1470 |
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
|
1471 |
val T = devarify (const_typ thy case_const); |
40844 | 1472 |
val Ts = binder_types T; |
37438 | 1473 |
val T_cong = nth Ts pos; |
1474 |
fun mk_prem z = Free (z, T_cong); |
|
1475 |
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
|
1476 |
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
|
1477 |
in |
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
54604
diff
changeset
|
1478 |
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
|
1479 |
(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
|
1480 |
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
|
1481 |
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
|
1482 |
end; |
37438 | 1483 |
|
66251
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
1484 |
fun declare_case_global thm thy = |
34244 | 1485 |
let |
43634 | 1486 |
val (case_const, (k, cos)) = case_cert thm; |
66310
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1487 |
fun get_type_of_constr c = case get_type_of_constr_or_abstr thy c of |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1488 |
SOME (c, false) => SOME c |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1489 |
| _ => NONE; |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1490 |
val cos_with_tycos = |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1491 |
(map_filter o Option.map) (fn c => (c, get_type_of_constr c)) cos; |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1492 |
val _ = case map_filter (fn (c, NONE) => SOME c | _ => NONE) cos_with_tycos of |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1493 |
[] => () |
45430 | 1494 |
| cs => error ("Non-constructor(s) in case certificate: " ^ commas_quote cs); |
66310
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1495 |
val tycos = distinct (op =) (map_filter snd cos_with_tycos); |
75399
cdf84288d93c
pass constructor arity as part of case certficiate
haftmann
parents:
75353
diff
changeset
|
1496 |
val schema = (1 + Int.max (1, length cos), |
cdf84288d93c
pass constructor arity as part of case certficiate
haftmann
parents:
75353
diff
changeset
|
1497 |
(k, (map o Option.map) (fn c => (c, args_number thy c)) cos)); |
66310
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1498 |
val cong = case_cong thy case_const schema; |
37438 | 1499 |
in |
1500 |
thy |
|
66310
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1501 |
|> modify_specs (map_cases (History.register case_const |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1502 |
(Case {schema = schema, tycos = tycos, cong = cong})) |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1503 |
#> map_types (fold (fn tyco => History.modify_entry tyco |
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1504 |
(add_case_combinator case_const)) tycos)) |
37438 | 1505 |
end; |
34244 | 1506 |
|
66251
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
1507 |
fun declare_undefined_global c = |
66310
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1508 |
(modify_specs o map_cases) (History.register c Undefined); |
34244 | 1509 |
|
1510 |
||
66310
e8d2862ec203
simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents:
66260
diff
changeset
|
1511 |
(* attributes *) |
31998
2c7a24f74db9
code attributes use common underscore convention
haftmann
parents:
31971
diff
changeset
|
1512 |
|
66251
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
1513 |
fun code_attribute f = Thm.declaration_attribute |
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
1514 |
(fn thm => Context.mapping (f thm) I); |
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
1515 |
|
68773 | 1516 |
fun code_thm_attribute g f = |
73968 | 1517 |
Scan.lift (g |-- Scan.succeed (code_attribute f)); |
68773 | 1518 |
|
1519 |
fun code_const_attribute g f = |
|
73968 | 1520 |
Scan.lift (g -- Args.colon) |-- Scan.repeat1 Args.term |
1521 |
>> (fn ts => code_attribute (K (fold (fn t => fn thy => f ((check_const thy o Logic.unvarify_types_global) t) thy) ts))); |
|
66251
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
1522 |
|
53171 | 1523 |
val _ = Theory.setup |
31962 | 1524 |
(let |
31998
2c7a24f74db9
code attributes use common underscore convention
haftmann
parents:
31971
diff
changeset
|
1525 |
val code_attribute_parser = |
68773 | 1526 |
code_thm_attribute (Args.$$$ "equation") |
1527 |
(fn thm => generic_add_eqn Liberal (thm, true)) |
|
1528 |
|| code_thm_attribute (Args.$$$ "nbe") |
|
1529 |
(fn thm => generic_add_eqn Liberal (thm, false)) |
|
1530 |
|| code_thm_attribute (Args.$$$ "abstract") |
|
1531 |
(generic_declare_abstract_eqn Liberal) |
|
1532 |
|| code_thm_attribute (Args.$$$ "abstype") |
|
1533 |
(generic_declare_abstype Liberal) |
|
1534 |
|| code_thm_attribute Args.del |
|
1535 |
del_eqn_global |
|
1536 |
|| code_const_attribute (Args.$$$ "abort") |
|
1537 |
declare_aborting_global |
|
1538 |
|| code_const_attribute (Args.$$$ "drop") |
|
1539 |
declare_unimplemented_global |
|
66251
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
1540 |
|| Scan.succeed (code_attribute |
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66189
diff
changeset
|
1541 |
add_maybe_abs_eqn_liberal); |
31962 | 1542 |
in |
73968 | 1543 |
Attrib.setup \<^binding>\<open>code\<close> code_attribute_parser |
31998
2c7a24f74db9
code attributes use common underscore convention
haftmann
parents:
31971
diff
changeset
|
1544 |
"declare theorems for code generation" |
53171 | 1545 |
end); |
31962 | 1546 |
|
24219 | 1547 |
end; (*struct*) |
1548 |
||
1549 |
||
35226 | 1550 |
(* type-safe interfaces for data dependent on executable code *) |
24219 | 1551 |
|
34173
458ced35abb8
reduced code generator cache to the baremost minimum
haftmann
parents:
33977
diff
changeset
|
1552 |
functor Code_Data(Data: CODE_DATA_ARGS): CODE_DATA = |
24219 | 1553 |
struct |
1554 |
||
1555 |
type T = Data.T; |
|
1556 |
exception Data of T; |
|
1557 |
fun dest (Data x) = x |
|
1558 |
||
34173
458ced35abb8
reduced code generator cache to the baremost minimum
haftmann
parents:
33977
diff
changeset
|
1559 |
val kind = Code.declare_data (Data Data.empty); |
24219 | 1560 |
|
1561 |
val data_op = (kind, Data, dest); |
|
1562 |
||
39397 | 1563 |
fun change_yield (SOME thy) f = Code.change_yield_data data_op thy f |
1564 |
| change_yield NONE f = f Data.empty |
|
1565 |
||
1566 |
fun change some_thy f = snd (change_yield some_thy (pair () o f)); |
|
24219 | 1567 |
|
1568 |
end; |
|
1569 |
||
28143 | 1570 |
structure Code : CODE = struct open Code; end; |