author | wenzelm |
Thu, 07 Aug 2008 19:21:41 +0200 | |
changeset 27779 | 4569003b8813 |
parent 27610 | 8882d47e075f |
child 28015 | 11635f41abc1 |
permissions | -rw-r--r-- |
24219 | 1 |
(* Title: Pure/Isar/code_unit.ML |
2 |
ID: $Id$ |
|
3 |
Author: Florian Haftmann, TU Muenchen |
|
4 |
||
24844
98c006a30218
certificates for code generator case expressions
haftmann
parents:
24707
diff
changeset
|
5 |
Basic notions of code generation. Auxiliary. |
24219 | 6 |
*) |
7 |
||
8 |
signature CODE_UNIT = |
|
9 |
sig |
|
24844
98c006a30218
certificates for code generator case expressions
haftmann
parents:
24707
diff
changeset
|
10 |
(*generic non-sense*) |
98c006a30218
certificates for code generator case expressions
haftmann
parents:
24707
diff
changeset
|
11 |
val bad_thm: string -> 'a |
98c006a30218
certificates for code generator case expressions
haftmann
parents:
24707
diff
changeset
|
12 |
val error_thm: (thm -> thm) -> thm -> thm |
98c006a30218
certificates for code generator case expressions
haftmann
parents:
24707
diff
changeset
|
13 |
val warning_thm: (thm -> thm) -> thm -> thm option |
98c006a30218
certificates for code generator case expressions
haftmann
parents:
24707
diff
changeset
|
14 |
val try_thm: (thm -> thm) -> thm -> thm option |
98c006a30218
certificates for code generator case expressions
haftmann
parents:
24707
diff
changeset
|
15 |
|
98c006a30218
certificates for code generator case expressions
haftmann
parents:
24707
diff
changeset
|
16 |
(*typ instantiations*) |
26970 | 17 |
val typscheme: theory -> string * typ -> (string * sort) list * typ |
24844
98c006a30218
certificates for code generator case expressions
haftmann
parents:
24707
diff
changeset
|
18 |
val inst_thm: sort Vartab.table -> thm -> thm |
25540 | 19 |
val constrain_thm: sort -> thm -> thm |
24844
98c006a30218
certificates for code generator case expressions
haftmann
parents:
24707
diff
changeset
|
20 |
|
26747
f32fa5f5bdd1
moved 'trivial classes' to foundation of code generator
haftmann
parents:
26618
diff
changeset
|
21 |
(*constant aliasses*) |
26354 | 22 |
val add_const_alias: thm -> theory -> theory |
26610 | 23 |
val subst_alias: theory -> string -> string |
26747
f32fa5f5bdd1
moved 'trivial classes' to foundation of code generator
haftmann
parents:
26618
diff
changeset
|
24 |
val resubst_alias: theory -> string -> string |
f32fa5f5bdd1
moved 'trivial classes' to foundation of code generator
haftmann
parents:
26618
diff
changeset
|
25 |
val triv_classes: theory -> class list |
f32fa5f5bdd1
moved 'trivial classes' to foundation of code generator
haftmann
parents:
26618
diff
changeset
|
26 |
|
f32fa5f5bdd1
moved 'trivial classes' to foundation of code generator
haftmann
parents:
26618
diff
changeset
|
27 |
(*constants*) |
24219 | 28 |
val string_of_typ: theory -> typ -> string |
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24219
diff
changeset
|
29 |
val string_of_const: theory -> string -> string |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24219
diff
changeset
|
30 |
val no_args: theory -> string -> int |
26112 | 31 |
val check_const: theory -> term -> string |
24219 | 32 |
val read_bare_const: theory -> string -> string * typ |
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24219
diff
changeset
|
33 |
val read_const: theory -> string -> string |
24219 | 34 |
|
24844
98c006a30218
certificates for code generator case expressions
haftmann
parents:
24707
diff
changeset
|
35 |
(*constructor sets*) |
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24219
diff
changeset
|
36 |
val constrset_of_consts: theory -> (string * typ) list |
24219 | 37 |
-> string * ((string * sort) list * (string * typ list) list) |
38 |
||
24844
98c006a30218
certificates for code generator case expressions
haftmann
parents:
24707
diff
changeset
|
39 |
(*defining equations*) |
24219 | 40 |
val assert_rew: thm -> thm |
41 |
val mk_rew: thm -> thm |
|
42 |
val mk_func: thm -> thm |
|
26970 | 43 |
val head_func: thm -> string * ((string * sort) list * typ) |
24219 | 44 |
val expand_eta: int -> thm -> thm |
27582 | 45 |
val rewrite_func: simpset -> thm -> thm |
46 |
val rewrite_head: thm list -> thm -> thm |
|
24219 | 47 |
val norm_args: thm list -> thm list |
24844
98c006a30218
certificates for code generator case expressions
haftmann
parents:
24707
diff
changeset
|
48 |
val norm_varnames: (string -> string) -> (string -> string) -> thm list -> thm list |
98c006a30218
certificates for code generator case expressions
haftmann
parents:
24707
diff
changeset
|
49 |
|
98c006a30218
certificates for code generator case expressions
haftmann
parents:
24707
diff
changeset
|
50 |
(*case certificates*) |
98c006a30218
certificates for code generator case expressions
haftmann
parents:
24707
diff
changeset
|
51 |
val case_cert: thm -> string * (int * string list) |
24219 | 52 |
end; |
53 |
||
54 |
structure CodeUnit: CODE_UNIT = |
|
55 |
struct |
|
56 |
||
57 |
||
58 |
(* auxiliary *) |
|
59 |
||
60 |
exception BAD_THM of string; |
|
61 |
fun bad_thm msg = raise BAD_THM msg; |
|
62 |
fun error_thm f thm = f thm handle BAD_THM msg => error msg; |
|
63 |
fun warning_thm f thm = SOME (f thm) handle BAD_THM msg |
|
64 |
=> (warning ("code generator: " ^ msg); NONE); |
|
24624
b8383b1bbae3
distinction between regular and default code theorems
haftmann
parents:
24423
diff
changeset
|
65 |
fun try_thm f thm = SOME (f thm) handle BAD_THM _ => NONE; |
24219 | 66 |
|
26939
1035c89b4c02
moved global pretty/string_of functions from Sign to Syntax;
wenzelm
parents:
26747
diff
changeset
|
67 |
fun string_of_typ thy = setmp show_sorts true (Syntax.string_of_typ_global thy); |
25597
34860182b250
moved instance parameter management from class.ML to axclass.ML
haftmann
parents:
25540
diff
changeset
|
68 |
fun string_of_const thy c = case AxClass.inst_of_param thy c |
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24219
diff
changeset
|
69 |
of SOME (c, tyco) => Sign.extern_const thy c ^ " " ^ enclose "[" "]" (Sign.extern_type thy tyco) |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24219
diff
changeset
|
70 |
| NONE => Sign.extern_const thy c; |
24219 | 71 |
|
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24219
diff
changeset
|
72 |
fun no_args thy = length o fst o strip_type o Sign.the_const_type thy; |
24219 | 73 |
|
26354 | 74 |
|
75 |
(* utilities *) |
|
76 |
||
26970 | 77 |
fun typscheme thy (c, ty) = |
78 |
let |
|
79 |
fun dest (TVar ((v, 0), sort)) = (v, sort) |
|
80 |
| dest ty = error ("Illegal type parameter in type scheme: " ^ Syntax.string_of_typ_global thy ty); |
|
81 |
val vs = map dest (Sign.const_typargs thy (c, ty)); |
|
82 |
in (vs, ty) end; |
|
83 |
||
26354 | 84 |
fun inst_thm tvars' thm = |
85 |
let |
|
86 |
val thy = Thm.theory_of_thm thm; |
|
87 |
val tvars = (Term.add_tvars o Thm.prop_of) thm []; |
|
88 |
fun mk_inst (tvar as (v, _)) = case Vartab.lookup tvars' v |
|
89 |
of SOME sort => SOME (pairself (Thm.ctyp_of thy o TVar) (tvar, (v, sort))) |
|
90 |
| NONE => NONE; |
|
91 |
val insts = map_filter mk_inst tvars; |
|
92 |
in Thm.instantiate (insts, []) thm end; |
|
93 |
||
94 |
fun constrain_thm sort thm = |
|
95 |
let |
|
96 |
val thy = Thm.theory_of_thm thm; |
|
97 |
val constrain = curry (Sorts.inter_sort (Sign.classes_of thy)) sort |
|
98 |
val tvars = (Term.add_tvars o Thm.prop_of) thm []; |
|
99 |
fun mk_inst (tvar as (v, sort)) = pairself (Thm.ctyp_of thy o TVar o pair v) |
|
100 |
(sort, constrain sort) |
|
101 |
val insts = map mk_inst tvars; |
|
102 |
in Thm.instantiate (insts, []) thm end; |
|
103 |
||
104 |
fun expand_eta k thm = |
|
105 |
let |
|
106 |
val thy = Thm.theory_of_thm thm; |
|
107 |
val (lhs, rhs) = (Logic.dest_equals o Thm.plain_prop_of) thm; |
|
108 |
val (head, args) = strip_comb lhs; |
|
109 |
val l = if k = ~1 |
|
110 |
then (length o fst o strip_abs) rhs |
|
111 |
else Int.max (0, k - length args); |
|
112 |
val used = Name.make_context (map (fst o fst) (Term.add_vars lhs [])); |
|
113 |
fun get_name _ 0 = pair [] |
|
114 |
| get_name (Abs (v, ty, t)) k = |
|
115 |
Name.variants [v] |
|
116 |
##>> get_name t (k - 1) |
|
117 |
#>> (fn ([v'], vs') => (v', ty) :: vs') |
|
118 |
| get_name t k = |
|
119 |
let |
|
120 |
val (tys, _) = (strip_type o fastype_of) t |
|
121 |
in case tys |
|
122 |
of [] => raise TERM ("expand_eta", [t]) |
|
123 |
| ty :: _ => |
|
124 |
Name.variants [""] |
|
125 |
#-> (fn [v] => get_name (t $ Var ((v, 0), ty)) (k - 1) |
|
126 |
#>> (fn vs' => (v, ty) :: vs')) |
|
127 |
end; |
|
128 |
val (vs, _) = get_name rhs l used; |
|
129 |
fun expand (v, ty) thm = Drule.fun_cong_rule thm |
|
130 |
(Thm.cterm_of thy (Var ((v, 0), ty))); |
|
131 |
in |
|
132 |
thm |
|
133 |
|> fold expand vs |
|
134 |
|> Conv.fconv_rule Drule.beta_eta_conversion |
|
135 |
end; |
|
136 |
||
137 |
fun func_conv conv = |
|
138 |
let |
|
139 |
fun lhs_conv ct = if can Thm.dest_comb ct |
|
140 |
then (Conv.combination_conv lhs_conv conv) ct |
|
141 |
else Conv.all_conv ct; |
|
142 |
in Conv.combination_conv (Conv.arg_conv lhs_conv) conv end; |
|
143 |
||
144 |
fun head_conv conv = |
|
145 |
let |
|
146 |
fun lhs_conv ct = if can Thm.dest_comb ct |
|
147 |
then (Conv.fun_conv lhs_conv) ct |
|
148 |
else conv ct; |
|
149 |
in Conv.fun_conv (Conv.arg_conv lhs_conv) end; |
|
150 |
||
27582 | 151 |
val rewrite_func = Conv.fconv_rule o func_conv o Simplifier.rewrite; |
26354 | 152 |
val rewrite_head = Conv.fconv_rule o head_conv o MetaSimplifier.rewrite false; |
153 |
||
154 |
fun norm_args thms = |
|
155 |
let |
|
156 |
val num_args_of = length o snd o strip_comb o fst o Logic.dest_equals; |
|
157 |
val k = fold (curry Int.max o num_args_of o Thm.plain_prop_of) thms 0; |
|
158 |
in |
|
159 |
thms |
|
160 |
|> map (expand_eta k) |
|
161 |
|> map (Conv.fconv_rule Drule.beta_eta_conversion) |
|
162 |
end; |
|
163 |
||
164 |
fun canonical_tvars purify_tvar thm = |
|
165 |
let |
|
166 |
val ctyp = Thm.ctyp_of (Thm.theory_of_thm thm); |
|
167 |
fun tvars_subst_for thm = (fold_types o fold_atyps) |
|
168 |
(fn TVar (v_i as (v, _), sort) => let |
|
169 |
val v' = purify_tvar v |
|
170 |
in if v = v' then I |
|
171 |
else insert (op =) (v_i, (v', sort)) end |
|
172 |
| _ => I) (prop_of thm) []; |
|
173 |
fun mk_inst (v_i, (v', sort)) (maxidx, acc) = |
|
174 |
let |
|
175 |
val ty = TVar (v_i, sort) |
|
176 |
in |
|
177 |
(maxidx + 1, (ctyp ty, ctyp (TVar ((v', maxidx), sort))) :: acc) |
|
178 |
end; |
|
179 |
val maxidx = Thm.maxidx_of thm + 1; |
|
180 |
val (_, inst) = fold mk_inst (tvars_subst_for thm) (maxidx + 1, []); |
|
181 |
in Thm.instantiate (inst, []) thm end; |
|
182 |
||
183 |
fun canonical_vars purify_var thm = |
|
184 |
let |
|
185 |
val cterm = Thm.cterm_of (Thm.theory_of_thm thm); |
|
186 |
fun vars_subst_for thm = fold_aterms |
|
187 |
(fn Var (v_i as (v, _), ty) => let |
|
188 |
val v' = purify_var v |
|
189 |
in if v = v' then I |
|
190 |
else insert (op =) (v_i, (v', ty)) end |
|
191 |
| _ => I) (prop_of thm) []; |
|
192 |
fun mk_inst (v_i as (v, i), (v', ty)) (maxidx, acc) = |
|
193 |
let |
|
194 |
val t = Var (v_i, ty) |
|
195 |
in |
|
196 |
(maxidx + 1, (cterm t, cterm (Var ((v', maxidx), ty))) :: acc) |
|
197 |
end; |
|
198 |
val maxidx = Thm.maxidx_of thm + 1; |
|
199 |
val (_, inst) = fold mk_inst (vars_subst_for thm) (maxidx + 1, []); |
|
200 |
in Thm.instantiate ([], inst) thm end; |
|
201 |
||
202 |
fun canonical_absvars purify_var thm = |
|
203 |
let |
|
204 |
val t = Thm.plain_prop_of thm; |
|
205 |
val t' = Term.map_abs_vars purify_var t; |
|
206 |
in Thm.rename_boundvars t t' thm end; |
|
207 |
||
208 |
fun norm_varnames purify_tvar purify_var thms = |
|
209 |
let |
|
210 |
fun burrow_thms f [] = [] |
|
211 |
| burrow_thms f thms = |
|
212 |
thms |
|
213 |
|> Conjunction.intr_balanced |
|
214 |
|> f |
|
215 |
|> Conjunction.elim_balanced (length thms) |
|
216 |
in |
|
217 |
thms |
|
218 |
|> burrow_thms (canonical_tvars purify_tvar) |
|
219 |
|> map (canonical_vars purify_var) |
|
220 |
|> map (canonical_absvars purify_var) |
|
221 |
|> map Drule.zero_var_indexes |
|
222 |
end; |
|
223 |
||
224 |
(* const aliasses *) |
|
225 |
||
226 |
structure ConstAlias = TheoryDataFun |
|
227 |
( |
|
26747
f32fa5f5bdd1
moved 'trivial classes' to foundation of code generator
haftmann
parents:
26618
diff
changeset
|
228 |
type T = ((string * string) * thm) list * class list; |
f32fa5f5bdd1
moved 'trivial classes' to foundation of code generator
haftmann
parents:
26618
diff
changeset
|
229 |
val empty = ([], []); |
26354 | 230 |
val copy = I; |
26618 | 231 |
val extend = I; |
26747
f32fa5f5bdd1
moved 'trivial classes' to foundation of code generator
haftmann
parents:
26618
diff
changeset
|
232 |
fun merge _ ((alias1, classes1), (alias2, classes2)) = |
f32fa5f5bdd1
moved 'trivial classes' to foundation of code generator
haftmann
parents:
26618
diff
changeset
|
233 |
(Library.merge (eq_snd Thm.eq_thm_prop) (alias1, alias2), |
f32fa5f5bdd1
moved 'trivial classes' to foundation of code generator
haftmann
parents:
26618
diff
changeset
|
234 |
Library.merge (op =) (classes1, classes2)); |
26354 | 235 |
); |
236 |
||
237 |
fun add_const_alias thm = |
|
238 |
let |
|
239 |
val t = Thm.prop_of thm; |
|
240 |
val thy = Thm.theory_of_thm thm; |
|
241 |
val lhs_rhs = case try Logic.dest_equals t |
|
242 |
of SOME lhs_rhs => lhs_rhs |
|
243 |
| _ => error ("Not an equation: " ^ Display.string_of_thm thm); |
|
244 |
val c_c' = case try (pairself (AxClass.unoverload_const thy o dest_Const)) lhs_rhs |
|
245 |
of SOME c_c' => c_c' |
|
246 |
| _ => error ("Not an equation with two constants: " ^ Display.string_of_thm thm); |
|
26747
f32fa5f5bdd1
moved 'trivial classes' to foundation of code generator
haftmann
parents:
26618
diff
changeset
|
247 |
val some_class = the_list (AxClass.class_of_param thy (snd c_c')); |
f32fa5f5bdd1
moved 'trivial classes' to foundation of code generator
haftmann
parents:
26618
diff
changeset
|
248 |
in |
f32fa5f5bdd1
moved 'trivial classes' to foundation of code generator
haftmann
parents:
26618
diff
changeset
|
249 |
ConstAlias.map (fn (alias, classes) => |
f32fa5f5bdd1
moved 'trivial classes' to foundation of code generator
haftmann
parents:
26618
diff
changeset
|
250 |
((c_c', thm) :: alias, fold (insert (op =)) some_class classes)) |
f32fa5f5bdd1
moved 'trivial classes' to foundation of code generator
haftmann
parents:
26618
diff
changeset
|
251 |
end; |
26354 | 252 |
|
253 |
fun rew_alias thm = |
|
254 |
let |
|
255 |
val thy = Thm.theory_of_thm thm; |
|
26747
f32fa5f5bdd1
moved 'trivial classes' to foundation of code generator
haftmann
parents:
26618
diff
changeset
|
256 |
in rewrite_head ((map snd o fst o ConstAlias.get) thy) thm end; |
26354 | 257 |
|
258 |
fun subst_alias thy c = ConstAlias.get thy |
|
26747
f32fa5f5bdd1
moved 'trivial classes' to foundation of code generator
haftmann
parents:
26618
diff
changeset
|
259 |
|> fst |
26354 | 260 |
|> get_first (fn ((c', c''), _) => if c = c' then SOME c'' else NONE) |
261 |
|> the_default c; |
|
262 |
||
26747
f32fa5f5bdd1
moved 'trivial classes' to foundation of code generator
haftmann
parents:
26618
diff
changeset
|
263 |
fun resubst_alias thy = |
f32fa5f5bdd1
moved 'trivial classes' to foundation of code generator
haftmann
parents:
26618
diff
changeset
|
264 |
let |
f32fa5f5bdd1
moved 'trivial classes' to foundation of code generator
haftmann
parents:
26618
diff
changeset
|
265 |
val alias = fst (ConstAlias.get thy); |
f32fa5f5bdd1
moved 'trivial classes' to foundation of code generator
haftmann
parents:
26618
diff
changeset
|
266 |
val subst_inst_param = Option.map fst o AxClass.inst_of_param thy; |
f32fa5f5bdd1
moved 'trivial classes' to foundation of code generator
haftmann
parents:
26618
diff
changeset
|
267 |
fun subst_alias c = |
f32fa5f5bdd1
moved 'trivial classes' to foundation of code generator
haftmann
parents:
26618
diff
changeset
|
268 |
get_first (fn ((c', c''), _) => if c = c'' then SOME c' else NONE) alias; |
f32fa5f5bdd1
moved 'trivial classes' to foundation of code generator
haftmann
parents:
26618
diff
changeset
|
269 |
in |
f32fa5f5bdd1
moved 'trivial classes' to foundation of code generator
haftmann
parents:
26618
diff
changeset
|
270 |
perhaps subst_inst_param |
f32fa5f5bdd1
moved 'trivial classes' to foundation of code generator
haftmann
parents:
26618
diff
changeset
|
271 |
#> perhaps subst_alias |
f32fa5f5bdd1
moved 'trivial classes' to foundation of code generator
haftmann
parents:
26618
diff
changeset
|
272 |
end; |
f32fa5f5bdd1
moved 'trivial classes' to foundation of code generator
haftmann
parents:
26618
diff
changeset
|
273 |
|
f32fa5f5bdd1
moved 'trivial classes' to foundation of code generator
haftmann
parents:
26618
diff
changeset
|
274 |
val triv_classes = snd o ConstAlias.get; |
f32fa5f5bdd1
moved 'trivial classes' to foundation of code generator
haftmann
parents:
26618
diff
changeset
|
275 |
|
26610 | 276 |
(* reading constants as terms *) |
24219 | 277 |
|
26112 | 278 |
fun check_bare_const thy t = case try dest_Const t |
279 |
of SOME c_ty => c_ty |
|
26939
1035c89b4c02
moved global pretty/string_of functions from Sign to Syntax;
wenzelm
parents:
26747
diff
changeset
|
280 |
| NONE => error ("Not a constant: " ^ Syntax.string_of_term_global thy t); |
26112 | 281 |
|
26519 | 282 |
fun check_const thy = subst_alias thy o AxClass.unoverload_const thy o apfst (subst_alias thy) |
283 |
o check_bare_const thy; |
|
26112 | 284 |
|
285 |
fun read_bare_const thy = check_bare_const thy o Syntax.read_term_global thy; |
|
24219 | 286 |
|
26519 | 287 |
fun read_const thy = subst_alias thy o AxClass.unoverload_const thy o apfst (subst_alias thy) |
288 |
o read_bare_const thy; |
|
24219 | 289 |
|
290 |
||
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24219
diff
changeset
|
291 |
(* constructor sets *) |
24219 | 292 |
|
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24219
diff
changeset
|
293 |
fun constrset_of_consts thy cs = |
24219 | 294 |
let |
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24219
diff
changeset
|
295 |
fun no_constr (c, ty) = error ("Not a datatype constructor: " ^ string_of_const thy c |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24219
diff
changeset
|
296 |
^ " :: " ^ string_of_typ thy ty); |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24219
diff
changeset
|
297 |
fun last_typ c_ty ty = |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24219
diff
changeset
|
298 |
let |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24219
diff
changeset
|
299 |
val frees = typ_tfrees ty; |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24219
diff
changeset
|
300 |
val (tyco, vs) = ((apsnd o map) (dest_TFree) o dest_Type o snd o strip_type) ty |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24219
diff
changeset
|
301 |
handle TYPE _ => no_constr c_ty |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24219
diff
changeset
|
302 |
val _ = if has_duplicates (eq_fst (op =)) vs then no_constr c_ty else (); |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24219
diff
changeset
|
303 |
val _ = if length frees <> length vs then no_constr c_ty else (); |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24219
diff
changeset
|
304 |
in (tyco, vs) end; |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24219
diff
changeset
|
305 |
fun ty_sorts (c, ty) = |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24219
diff
changeset
|
306 |
let |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24219
diff
changeset
|
307 |
val ty_decl = (Logic.unvarifyT o Sign.the_const_type thy) c; |
26970 | 308 |
val (tyco, _) = last_typ (c, ty) ty_decl; |
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24219
diff
changeset
|
309 |
val (_, vs) = last_typ (c, ty) ty; |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24219
diff
changeset
|
310 |
in ((tyco, map snd vs), (c, (map fst vs, ty_decl))) end; |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24219
diff
changeset
|
311 |
fun add ((tyco', sorts'), c) ((tyco, sorts), cs) = |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24219
diff
changeset
|
312 |
let |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24219
diff
changeset
|
313 |
val _ = if tyco' <> tyco |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24219
diff
changeset
|
314 |
then error "Different type constructors in constructor set" |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24219
diff
changeset
|
315 |
else (); |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24219
diff
changeset
|
316 |
val sorts'' = map2 (curry (Sorts.inter_sort (Sign.classes_of thy))) sorts' sorts |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24219
diff
changeset
|
317 |
in ((tyco, sorts), c :: cs) end; |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24219
diff
changeset
|
318 |
fun inst vs' (c, (vs, ty)) = |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24219
diff
changeset
|
319 |
let |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24219
diff
changeset
|
320 |
val the_v = the o AList.lookup (op =) (vs ~~ vs'); |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24219
diff
changeset
|
321 |
val ty' = map_atyps (fn TFree (v, _) => TFree (the_v v)) ty; |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24219
diff
changeset
|
322 |
in (c, (fst o strip_type) ty') end; |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24219
diff
changeset
|
323 |
val c' :: cs' = map ty_sorts cs; |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24219
diff
changeset
|
324 |
val ((tyco, sorts), cs'') = fold add cs' (apsnd single c'); |
24848 | 325 |
val vs = Name.names Name.context Name.aT sorts; |
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24219
diff
changeset
|
326 |
val cs''' = map (inst vs) cs''; |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24219
diff
changeset
|
327 |
in (tyco, (vs, cs''')) end; |
24219 | 328 |
|
329 |
||
26239
e105d24d15c1
some steps towards a refined treatment of equality
haftmann
parents:
26112
diff
changeset
|
330 |
(* rewrite theorems *) |
24219 | 331 |
|
332 |
fun assert_rew thm = |
|
333 |
let |
|
334 |
val (lhs, rhs) = (Logic.dest_equals o Thm.plain_prop_of) thm |
|
335 |
handle TERM _ => bad_thm ("Not an equation: " ^ Display.string_of_thm thm) |
|
336 |
| THM _ => bad_thm ("Not an equation: " ^ Display.string_of_thm thm); |
|
337 |
fun vars_of t = fold_aterms |
|
338 |
(fn Var (v, _) => insert (op =) v |
|
339 |
| Free _ => bad_thm ("Illegal free variable in rewrite theorem\n" |
|
340 |
^ Display.string_of_thm thm) |
|
341 |
| _ => I) t []; |
|
342 |
fun tvars_of t = fold_term_types |
|
343 |
(fn _ => fold_atyps (fn TVar (v, _) => insert (op =) v |
|
344 |
| TFree _ => bad_thm |
|
345 |
("Illegal free type variable in rewrite theorem\n" ^ Display.string_of_thm thm))) t []; |
|
346 |
val lhs_vs = vars_of lhs; |
|
347 |
val rhs_vs = vars_of rhs; |
|
348 |
val lhs_tvs = tvars_of lhs; |
|
349 |
val rhs_tvs = tvars_of lhs; |
|
350 |
val _ = if null (subtract (op =) lhs_vs rhs_vs) |
|
351 |
then () |
|
352 |
else bad_thm ("Free variables on right hand side of rewrite theorem\n" |
|
353 |
^ Display.string_of_thm thm); |
|
354 |
val _ = if null (subtract (op =) lhs_tvs rhs_tvs) |
|
355 |
then () |
|
356 |
else bad_thm ("Free type variables on right hand side of rewrite theorem\n" |
|
357 |
^ Display.string_of_thm thm) |
|
358 |
in thm end; |
|
359 |
||
360 |
fun mk_rew thm = |
|
361 |
let |
|
362 |
val thy = Thm.theory_of_thm thm; |
|
363 |
val ctxt = ProofContext.init thy; |
|
364 |
in |
|
365 |
thm |
|
366 |
|> LocalDefs.meta_rewrite_rule ctxt |
|
367 |
|> assert_rew |
|
368 |
end; |
|
369 |
||
370 |
||
26239
e105d24d15c1
some steps towards a refined treatment of equality
haftmann
parents:
26112
diff
changeset
|
371 |
(* defining equations *) |
e105d24d15c1
some steps towards a refined treatment of equality
haftmann
parents:
26112
diff
changeset
|
372 |
|
24219 | 373 |
fun assert_func thm = |
374 |
let |
|
375 |
val thy = Thm.theory_of_thm thm; |
|
376 |
val (head, args) = (strip_comb o fst o Logic.dest_equals |
|
377 |
o ObjectLogic.drop_judgment thy o Thm.plain_prop_of) thm; |
|
378 |
val _ = case head of Const _ => () | _ => |
|
379 |
bad_thm ("Equation not headed by constant\n" ^ Display.string_of_thm thm); |
|
380 |
val _ = |
|
381 |
if has_duplicates (op =) |
|
382 |
((fold o fold_aterms) (fn Var (v, _) => cons v |
|
383 |
| _ => I |
|
384 |
) args []) |
|
385 |
then bad_thm ("Duplicated variables on left hand side of equation\n" |
|
386 |
^ Display.string_of_thm thm) |
|
387 |
else () |
|
388 |
fun check _ (Abs _) = bad_thm |
|
389 |
("Abstraction on left hand side of equation\n" |
|
390 |
^ Display.string_of_thm thm) |
|
391 |
| check 0 (Var _) = () |
|
392 |
| check _ (Var _) = bad_thm |
|
393 |
("Variable with application on left hand side of defining equation\n" |
|
394 |
^ Display.string_of_thm thm) |
|
395 |
| check n (t1 $ t2) = (check (n+1) t1; check 0 t2) |
|
396 |
| check n (Const (_, ty)) = if n <> (length o fst o strip_type) ty |
|
397 |
then bad_thm |
|
398 |
("Partially applied constant on left hand side of equation\n" |
|
399 |
^ Display.string_of_thm thm) |
|
400 |
else (); |
|
401 |
val _ = map (check 0) args; |
|
402 |
in thm end; |
|
403 |
||
26354 | 404 |
val mk_func = rew_alias o assert_func o mk_rew; |
24219 | 405 |
|
406 |
fun head_func thm = |
|
407 |
let |
|
408 |
val thy = Thm.theory_of_thm thm; |
|
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24219
diff
changeset
|
409 |
val Const (c, ty) = (fst o strip_comb o fst o Logic.dest_equals |
27610 | 410 |
o Thm.plain_prop_of) thm; |
26970 | 411 |
in (c, typscheme thy (c, ty)) end; |
24219 | 412 |
|
413 |
||
24844
98c006a30218
certificates for code generator case expressions
haftmann
parents:
24707
diff
changeset
|
414 |
(* case cerificates *) |
98c006a30218
certificates for code generator case expressions
haftmann
parents:
24707
diff
changeset
|
415 |
|
24917 | 416 |
fun case_certificate thm = |
24844
98c006a30218
certificates for code generator case expressions
haftmann
parents:
24707
diff
changeset
|
417 |
let |
98c006a30218
certificates for code generator case expressions
haftmann
parents:
24707
diff
changeset
|
418 |
val thy = Thm.theory_of_thm thm; |
24917 | 419 |
val ((head, raw_case_expr), cases) = (apfst Logic.dest_equals |
420 |
o apsnd Logic.dest_conjunctions o Logic.dest_implies o Thm.prop_of) thm; |
|
24844
98c006a30218
certificates for code generator case expressions
haftmann
parents:
24707
diff
changeset
|
421 |
val _ = case head of Free _ => true |
98c006a30218
certificates for code generator case expressions
haftmann
parents:
24707
diff
changeset
|
422 |
| Var _ => true |
98c006a30218
certificates for code generator case expressions
haftmann
parents:
24707
diff
changeset
|
423 |
| _ => raise TERM ("case_cert", []); |
24917 | 424 |
val ([(case_var, _)], case_expr) = Term.strip_abs_eta 1 raw_case_expr; |
425 |
val (Const (case_const, _), raw_params) = strip_comb case_expr; |
|
426 |
val n = find_index (fn Free (v, _) => v = case_var | _ => false) raw_params; |
|
427 |
val _ = if n = ~1 then raise TERM ("case_cert", []) else (); |
|
428 |
val params = map (fst o dest_Var) (nth_drop n raw_params); |
|
429 |
fun dest_case t = |
|
24844
98c006a30218
certificates for code generator case expressions
haftmann
parents:
24707
diff
changeset
|
430 |
let |
98c006a30218
certificates for code generator case expressions
haftmann
parents:
24707
diff
changeset
|
431 |
val (head' $ t_co, rhs) = Logic.dest_equals t; |
98c006a30218
certificates for code generator case expressions
haftmann
parents:
24707
diff
changeset
|
432 |
val _ = if head' = head then () else raise TERM ("case_cert", []); |
98c006a30218
certificates for code generator case expressions
haftmann
parents:
24707
diff
changeset
|
433 |
val (Const (co, _), args) = strip_comb t_co; |
98c006a30218
certificates for code generator case expressions
haftmann
parents:
24707
diff
changeset
|
434 |
val (Var (param, _), args') = strip_comb rhs; |
98c006a30218
certificates for code generator case expressions
haftmann
parents:
24707
diff
changeset
|
435 |
val _ = if args' = args then () else raise TERM ("case_cert", []); |
98c006a30218
certificates for code generator case expressions
haftmann
parents:
24707
diff
changeset
|
436 |
in (param, co) end; |
24917 | 437 |
fun analyze_cases cases = |
24844
98c006a30218
certificates for code generator case expressions
haftmann
parents:
24707
diff
changeset
|
438 |
let |
24917 | 439 |
val co_list = fold (AList.update (op =) o dest_case) cases []; |
24844
98c006a30218
certificates for code generator case expressions
haftmann
parents:
24707
diff
changeset
|
440 |
in map (the o AList.lookup (op =) co_list) params end; |
98c006a30218
certificates for code generator case expressions
haftmann
parents:
24707
diff
changeset
|
441 |
fun analyze_let t = |
98c006a30218
certificates for code generator case expressions
haftmann
parents:
24707
diff
changeset
|
442 |
let |
98c006a30218
certificates for code generator case expressions
haftmann
parents:
24707
diff
changeset
|
443 |
val (head' $ arg, Var (param', _) $ arg') = Logic.dest_equals t; |
98c006a30218
certificates for code generator case expressions
haftmann
parents:
24707
diff
changeset
|
444 |
val _ = if head' = head then () else raise TERM ("case_cert", []); |
98c006a30218
certificates for code generator case expressions
haftmann
parents:
24707
diff
changeset
|
445 |
val _ = if arg' = arg then () else raise TERM ("case_cert", []); |
98c006a30218
certificates for code generator case expressions
haftmann
parents:
24707
diff
changeset
|
446 |
val _ = if [param'] = params then () else raise TERM ("case_cert", []); |
98c006a30218
certificates for code generator case expressions
haftmann
parents:
24707
diff
changeset
|
447 |
in [] end; |
24917 | 448 |
fun analyze (cases as [let_case]) = |
449 |
(analyze_cases cases handle Bind => analyze_let let_case) |
|
450 |
| analyze cases = analyze_cases cases; |
|
451 |
in (case_const, (n, analyze cases)) end; |
|
452 |
||
453 |
fun case_cert thm = case_certificate thm |
|
454 |
handle Bind => error "bad case certificate" |
|
455 |
| TERM _ => error "bad case certificate"; |
|
24844
98c006a30218
certificates for code generator case expressions
haftmann
parents:
24707
diff
changeset
|
456 |
|
24219 | 457 |
end; |