author | wenzelm |
Tue, 25 Sep 2007 13:28:37 +0200 | |
changeset 24707 | dfeb98f84e93 |
parent 24701 | f8bfd592a6dc |
child 24731 | c25aa6ae64ec |
permissions | -rw-r--r-- |
24218 | 1 |
(* Title: Pure/Isar/class.ML |
2 |
ID: $Id$ |
|
3 |
Author: Florian Haftmann, TU Muenchen |
|
4 |
||
5 |
Type classes derived from primitive axclasses and locales. |
|
6 |
*) |
|
7 |
||
8 |
signature CLASS = |
|
9 |
sig |
|
10 |
val fork_mixfix: bool -> string option -> mixfix -> mixfix * mixfix |
|
11 |
||
12 |
val axclass_cmd: bstring * xstring list |
|
24589 | 13 |
-> ((bstring * Attrib.src list) * string list) list |
14 |
-> theory -> class * theory |
|
15 |
val classrel_cmd: xstring * xstring -> theory -> Proof.state |
|
24218 | 16 |
val class: bstring -> class list -> Element.context_i Locale.element list |
17 |
-> string list -> theory -> string * Proof.context |
|
24589 | 18 |
val class_cmd: bstring -> xstring list -> Element.context Locale.element list |
19 |
-> xstring list -> theory -> string * Proof.context |
|
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
20 |
val add_const_in_class: string -> (string * term) * Syntax.mixfix |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
21 |
-> theory -> theory |
24589 | 22 |
val interpretation_in_class: class * class -> theory -> Proof.state |
23 |
val interpretation_in_class_cmd: xstring * xstring -> theory -> Proof.state |
|
24 |
val prove_interpretation_in_class: tactic -> class * class -> theory -> theory |
|
25 |
val intro_classes_tac: thm list -> tactic |
|
26 |
val default_intro_classes_tac: thm list -> tactic |
|
27 |
val class_of_locale: theory -> string -> class option |
|
28 |
val print_classes: theory -> unit |
|
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
29 |
|
24589 | 30 |
val instance_arity: (theory -> theory) -> arity list -> theory -> Proof.state |
31 |
val instance: arity list -> ((bstring * Attrib.src list) * term) list |
|
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
32 |
-> (thm list -> theory -> theory) |
24218 | 33 |
-> theory -> Proof.state |
24589 | 34 |
val instance_cmd: (bstring * xstring list * xstring) list |
35 |
-> ((bstring * Attrib.src list) * xstring) list |
|
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
36 |
-> (thm list -> theory -> theory) |
24218 | 37 |
-> theory -> Proof.state |
24589 | 38 |
val prove_instance: tactic -> arity list |
24218 | 39 |
-> ((bstring * Attrib.src list) * term) list |
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
40 |
-> theory -> thm list * theory |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
41 |
val unoverload: theory -> conv |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
42 |
val overload: theory -> conv |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
43 |
val unoverload_const: theory -> string * typ -> string |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
44 |
val inst_const: theory -> string * string -> string |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
45 |
val param_const: theory -> string -> (string * string) option |
24435 | 46 |
val params_of_sort: theory -> sort -> (string * (string * typ)) list |
24657 | 47 |
|
48 |
(*experimental*) |
|
49 |
val init_ref: (class -> Proof.context -> (theory -> theory) * Proof.context) ref |
|
50 |
val init: class -> Proof.context -> (theory -> theory) * Proof.context; |
|
51 |
val init_default: class -> Proof.context -> (theory -> theory) * Proof.context; |
|
52 |
val remove_constraints: class -> theory -> (string * typ) list * theory |
|
53 |
val class_term_check: theory -> class -> term list -> Proof.context -> term list * Proof.context |
|
54 |
val local_param: theory -> class -> string -> (term * (class * int)) option |
|
24218 | 55 |
end; |
56 |
||
57 |
structure Class : CLASS = |
|
58 |
struct |
|
59 |
||
60 |
(** auxiliary **) |
|
61 |
||
62 |
fun fork_mixfix is_loc some_class mx = |
|
63 |
let |
|
64 |
val mx' = Syntax.unlocalize_mixfix mx; |
|
65 |
val mx_global = if is_some some_class orelse (is_loc andalso mx = mx') |
|
66 |
then NoSyn else mx'; |
|
67 |
val mx_local = if is_loc then mx else NoSyn; |
|
68 |
in (mx_global, mx_local) end; |
|
69 |
||
24589 | 70 |
fun prove_interpretation tac prfx_atts expr insts = |
71 |
Locale.interpretation_i I prfx_atts expr insts |
|
72 |
#> Proof.global_terminal_proof |
|
73 |
(Method.Basic (K (Method.SIMPLE_METHOD tac), Position.none), NONE) |
|
74 |
#> ProofContext.theory_of; |
|
75 |
||
76 |
fun prove_interpretation_in tac after_qed (name, expr) = |
|
77 |
Locale.interpretation_in_locale |
|
78 |
(ProofContext.theory after_qed) (name, expr) |
|
79 |
#> Proof.global_terminal_proof |
|
80 |
(Method.Basic (K (Method.SIMPLE_METHOD tac), Position.none), NONE) |
|
81 |
#> ProofContext.theory_of; |
|
82 |
||
83 |
fun OF_LAST thm1 thm2 = |
|
84 |
let |
|
85 |
val n = (length o Logic.strip_imp_prems o prop_of) thm2; |
|
86 |
in (thm1 RSN (n, thm2)) end; |
|
87 |
||
88 |
fun strip_all_ofclass thy sort = |
|
89 |
let |
|
90 |
val typ = TVar ((AxClass.param_tyvarname, 0), sort); |
|
91 |
fun prem_inclass t = |
|
92 |
case Logic.strip_imp_prems t |
|
93 |
of ofcls :: _ => try Logic.dest_inclass ofcls |
|
94 |
| [] => NONE; |
|
95 |
fun strip_ofclass class thm = |
|
96 |
thm OF (fst o AxClass.of_sort thy (typ, [class])) AxClass.cache; |
|
97 |
fun strip thm = case (prem_inclass o Thm.prop_of) thm |
|
98 |
of SOME (_, class) => thm |> strip_ofclass class |> strip |
|
99 |
| NONE => thm; |
|
100 |
in strip end; |
|
101 |
||
24657 | 102 |
fun get_remove_contraint c thy = |
103 |
let |
|
104 |
val ty = Sign.the_const_constraint thy c; |
|
105 |
in |
|
106 |
thy |
|
107 |
|> Sign.add_const_constraint_i (c, NONE) |
|
108 |
|> pair (c, Logic.unvarifyT ty) |
|
109 |
end; |
|
110 |
||
24589 | 111 |
|
112 |
(** axclass command **) |
|
113 |
||
24218 | 114 |
fun axclass_cmd (class, raw_superclasses) raw_specs thy = |
115 |
let |
|
116 |
val ctxt = ProofContext.init thy; |
|
117 |
val superclasses = map (Sign.read_class thy) raw_superclasses; |
|
24589 | 118 |
val name_atts = map ((apsnd o map) (Attrib.attribute thy) o fst) |
119 |
raw_specs; |
|
120 |
val axiomss = ProofContext.read_propp (ctxt, map (map (rpair []) o snd) |
|
121 |
raw_specs) |
|
24218 | 122 |
|> snd |
123 |
|> (map o map) fst; |
|
24589 | 124 |
in |
125 |
AxClass.define_class (class, superclasses) [] |
|
126 |
(name_atts ~~ axiomss) thy |
|
127 |
end; |
|
24218 | 128 |
|
129 |
local |
|
130 |
||
131 |
fun gen_instance mk_prop add_thm after_qed insts thy = |
|
132 |
let |
|
133 |
fun after_qed' results = |
|
134 |
ProofContext.theory ((fold o fold) add_thm results #> after_qed); |
|
135 |
in |
|
136 |
thy |
|
137 |
|> ProofContext.init |
|
24589 | 138 |
|> Proof.theorem_i NONE after_qed' ((map (fn t => [(t, [])]) |
139 |
o maps (mk_prop thy)) insts) |
|
24218 | 140 |
end; |
141 |
||
142 |
in |
|
143 |
||
24589 | 144 |
val instance_arity = |
24218 | 145 |
gen_instance (Logic.mk_arities oo Sign.cert_arity) AxClass.add_arity; |
24589 | 146 |
val classrel = |
24218 | 147 |
gen_instance (single oo (Logic.mk_classrel oo AxClass.cert_classrel)) |
148 |
AxClass.add_classrel I o single; |
|
24589 | 149 |
val classrel_cmd = |
150 |
gen_instance (single oo (Logic.mk_classrel oo AxClass.read_classrel)) |
|
151 |
AxClass.add_classrel I o single; |
|
24218 | 152 |
|
153 |
end; (*local*) |
|
154 |
||
155 |
||
24589 | 156 |
(** explicit constants for overloaded definitions **) |
24304 | 157 |
|
158 |
structure InstData = TheoryDataFun |
|
159 |
( |
|
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
160 |
type T = (string * thm) Symtab.table Symtab.table * (string * string) Symtab.table; |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
161 |
(*constant name ~> type constructor ~> (constant name, equation), |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
162 |
constant name ~> (constant name, type constructor)*) |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
163 |
val empty = (Symtab.empty, Symtab.empty); |
24304 | 164 |
val copy = I; |
165 |
val extend = I; |
|
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
166 |
fun merge _ ((taba1, tabb1), (taba2, tabb2)) = |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
167 |
(Symtab.join (K (Symtab.merge (K true))) (taba1, taba2), |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
168 |
Symtab.merge (K true) (tabb1, tabb2)); |
24304 | 169 |
); |
170 |
||
24589 | 171 |
fun inst_thms f thy = (Symtab.fold (Symtab.fold (cons o f o snd o snd) o snd) o fst) |
172 |
(InstData.get thy) []; |
|
173 |
fun add_inst (c, tyco) inst = (InstData.map o apfst |
|
174 |
o Symtab.map_default (c, Symtab.empty)) (Symtab.update_new (tyco, inst)) |
|
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
175 |
#> (InstData.map o apsnd) (Symtab.update_new (fst inst, (c, tyco))); |
24304 | 176 |
|
24589 | 177 |
fun unoverload thy = MetaSimplifier.rewrite false (inst_thms I thy); |
178 |
fun overload thy = MetaSimplifier.rewrite false (inst_thms symmetric thy); |
|
24304 | 179 |
|
180 |
fun inst_const thy (c, tyco) = |
|
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
181 |
(fst o the o Symtab.lookup ((the o Symtab.lookup (fst (InstData.get thy))) c)) tyco; |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
182 |
fun unoverload_const thy (c_ty as (c, _)) = |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
183 |
case AxClass.class_of_param thy c |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
184 |
of SOME class => (case Sign.const_typargs thy c_ty |
24589 | 185 |
of [Type (tyco, _)] => (case Symtab.lookup |
186 |
((the o Symtab.lookup (fst (InstData.get thy))) c) tyco |
|
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
187 |
of SOME (c, _) => c |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
188 |
| NONE => c) |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
189 |
| [_] => c) |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
190 |
| NONE => c; |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
191 |
|
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
192 |
val param_const = Symtab.lookup o snd o InstData.get; |
24304 | 193 |
|
194 |
fun add_inst_def (class, tyco) (c, ty) thy = |
|
195 |
let |
|
196 |
val tyco_base = NameSpace.base tyco; |
|
197 |
val name_inst = NameSpace.base class ^ "_" ^ tyco_base ^ "_inst"; |
|
198 |
val c_inst_base = NameSpace.base c ^ "_" ^ tyco_base; |
|
199 |
in |
|
200 |
thy |
|
201 |
|> Sign.sticky_prefix name_inst |
|
24435 | 202 |
|> PureThy.simple_def ("", []) |
203 |
(((c_inst_base, ty, Syntax.NoSyn), []), Const (c, ty)) |
|
204 |
|-> (fn (c_inst, thm) => add_inst (c, tyco) (c_inst, symmetric thm)) |
|
205 |
|> Sign.restore_naming thy |
|
24304 | 206 |
end; |
207 |
||
208 |
fun add_inst_def' (class, tyco) (c, ty) thy = |
|
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
209 |
if case Symtab.lookup (fst (InstData.get thy)) c |
24304 | 210 |
of NONE => true |
211 |
| SOME tab => is_none (Symtab.lookup tab tyco) |
|
212 |
then add_inst_def (class, tyco) (c, Logic.unvarifyT ty) thy |
|
213 |
else thy; |
|
214 |
||
215 |
fun add_def ((class, tyco), ((name, prop), atts)) thy = |
|
216 |
let |
|
24589 | 217 |
val ((lhs as Const (c, ty), args), rhs) = |
218 |
(apfst Term.strip_comb o Logic.dest_equals) prop; |
|
24701 | 219 |
fun (*add_inst' def ([], (Const (c_inst, ty))) = |
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
220 |
if forall (fn TFree _ => true | _ => false) (Sign.const_typargs thy (c_inst, ty)) |
24304 | 221 |
then add_inst (c, tyco) (c_inst, def) |
222 |
else add_inst_def (class, tyco) (c, ty) |
|
24701 | 223 |
|*) add_inst' _ t = add_inst_def (class, tyco) (c, ty); |
24304 | 224 |
in |
225 |
thy |
|
226 |
|> PureThy.add_defs_i true [((name, prop), map (Attrib.attribute thy) atts)] |
|
227 |
|-> (fn [def] => add_inst' def (args, rhs) #> pair def) |
|
228 |
end; |
|
229 |
||
230 |
||
24589 | 231 |
(** instances with implicit parameter handling **) |
24218 | 232 |
|
233 |
local |
|
234 |
||
24707 | 235 |
fun gen_read_def thy prep_att parse_prop ((raw_name, raw_atts), raw_t) = |
24218 | 236 |
let |
24707 | 237 |
val ctxt = ProofContext.init thy; |
238 |
val t = parse_prop ctxt raw_t |> Syntax.check_prop ctxt; |
|
24218 | 239 |
val ((c, ty), _) = Sign.cert_def (Sign.pp thy) t; |
240 |
val atts = map (prep_att thy) raw_atts; |
|
241 |
val insts = Consts.typargs (Sign.consts_of thy) (c, ty); |
|
242 |
val name = case raw_name |
|
243 |
of "" => NONE |
|
244 |
| _ => SOME raw_name; |
|
245 |
in (c, (insts, ((name, t), atts))) end; |
|
246 |
||
24707 | 247 |
fun read_def_cmd thy = gen_read_def thy Attrib.intern_src Syntax.parse_prop; |
24218 | 248 |
fun read_def thy = gen_read_def thy (K I) (K I); |
249 |
||
24589 | 250 |
fun gen_instance prep_arity read_def do_proof raw_arities raw_defs after_qed theory = |
24218 | 251 |
let |
252 |
val arities = map (prep_arity theory) raw_arities; |
|
253 |
val _ = if null arities then error "at least one arity must be given" else (); |
|
254 |
val _ = case (duplicates (op =) o map #1) arities |
|
255 |
of [] => () |
|
256 |
| dupl_tycos => error ("type constructors occur more than once in arities: " |
|
257 |
^ (commas o map quote) dupl_tycos); |
|
24589 | 258 |
val super_sort = |
259 |
(Graph.all_succs o #classes o Sorts.rep_algebra o Sign.classes_of) theory; |
|
24218 | 260 |
fun get_consts_class tyco ty class = |
261 |
let |
|
262 |
val cs = (these o Option.map snd o try (AxClass.params_of_class theory)) class; |
|
263 |
val subst_ty = map_type_tfree (K ty); |
|
264 |
in |
|
24304 | 265 |
map (fn (c, ty) => (c, ((class, tyco), subst_ty ty))) cs |
24218 | 266 |
end; |
267 |
fun get_consts_sort (tyco, asorts, sort) = |
|
268 |
let |
|
24589 | 269 |
val ty = Type (tyco, map (fn (v, sort) => TVar ((v, 0), sort)) |
270 |
(Name.names Name.context "'a" asorts)) |
|
24218 | 271 |
in maps (get_consts_class tyco ty) (super_sort sort) end; |
272 |
val cs = maps get_consts_sort arities; |
|
273 |
fun mk_typnorm thy (ty, ty_sc) = |
|
274 |
case try (Sign.typ_match thy (Logic.varifyT ty_sc, ty)) Vartab.empty |
|
275 |
of SOME env => SOME (Logic.varifyT #> Envir.typ_subst_TVars env #> Logic.unvarifyT) |
|
276 |
| NONE => NONE; |
|
277 |
fun read_defs defs cs thy_read = |
|
278 |
let |
|
279 |
fun read raw_def cs = |
|
280 |
let |
|
281 |
val (c, (inst, ((name_opt, t), atts))) = read_def thy_read raw_def; |
|
282 |
val ty = Consts.instance (Sign.consts_of thy_read) (c, inst); |
|
24304 | 283 |
val ((class, tyco), ty') = case AList.lookup (op =) cs c |
24218 | 284 |
of NONE => error ("illegal definition for constant " ^ quote c) |
285 |
| SOME class_ty => class_ty; |
|
286 |
val name = case name_opt |
|
287 |
of NONE => Thm.def_name (Logic.name_arity (tyco, [], c)) |
|
288 |
| SOME name => name; |
|
289 |
val t' = case mk_typnorm thy_read (ty', ty) |
|
290 |
of NONE => error ("illegal definition for constant " ^ |
|
291 |
quote (c ^ "::" ^ setmp show_sorts true |
|
292 |
(Sign.string_of_typ thy_read) ty)) |
|
293 |
| SOME norm => map_types norm t |
|
294 |
in (((class, tyco), ((name, t'), atts)), AList.delete (op =) c cs) end; |
|
295 |
in fold_map read defs cs end; |
|
24304 | 296 |
val (defs, other_cs) = read_defs raw_defs cs |
24218 | 297 |
(fold Sign.primitive_arity arities (Theory.copy theory)); |
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
298 |
fun after_qed' cs defs = |
24304 | 299 |
fold Sign.add_const_constraint_i (map (apsnd SOME) cs) |
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
300 |
#> after_qed defs; |
24218 | 301 |
in |
302 |
theory |
|
303 |
|> fold_map get_remove_contraint (map fst cs |> distinct (op =)) |
|
24304 | 304 |
||>> fold_map add_def defs |
305 |
||> fold (fn (c, ((class, tyco), ty)) => add_inst_def' (class, tyco) (c, ty)) other_cs |
|
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
306 |
|-> (fn (cs, defs) => do_proof (after_qed' cs defs) arities defs) |
24218 | 307 |
end; |
308 |
||
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
309 |
fun tactic_proof tac f arities defs = |
24218 | 310 |
fold (fn arity => AxClass.prove_arity arity tac) arities |
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
311 |
#> f |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
312 |
#> pair defs; |
24218 | 313 |
|
314 |
in |
|
315 |
||
24589 | 316 |
val instance = gen_instance Sign.cert_arity read_def |
317 |
(fn f => fn arities => fn defs => instance_arity f arities); |
|
318 |
val instance_cmd = gen_instance Sign.read_arity read_def_cmd |
|
319 |
(fn f => fn arities => fn defs => instance_arity f arities); |
|
320 |
fun prove_instance tac arities defs = |
|
321 |
gen_instance Sign.cert_arity read_def |
|
322 |
(tactic_proof tac) arities defs (K I); |
|
24218 | 323 |
|
324 |
end; (*local*) |
|
325 |
||
326 |
||
327 |
||
24589 | 328 |
(** class data **) |
24218 | 329 |
|
330 |
datatype class_data = ClassData of { |
|
331 |
locale: string, |
|
332 |
consts: (string * string) list |
|
24657 | 333 |
(*locale parameter ~> theory constant name*), |
334 |
v: string, |
|
24589 | 335 |
inst: typ Symtab.table * term Symtab.table |
336 |
(*canonical interpretation*), |
|
24657 | 337 |
intro: thm, |
338 |
defs: thm list, |
|
339 |
localized: (string * (term * (class * int))) list |
|
340 |
(*theory constant name ~> (locale parameter, (class, instantiaton index of class typ))*) |
|
341 |
}; |
|
24218 | 342 |
|
24657 | 343 |
fun rep_class_data (ClassData d) = d; |
344 |
fun mk_class_data ((locale, consts, v, inst, intro), (defs, localized)) = |
|
345 |
ClassData { locale = locale, consts = consts, v = v, inst = inst, intro = intro, |
|
346 |
defs = defs, localized = localized }; |
|
347 |
fun map_class_data f (ClassData { locale, consts, v, inst, intro, defs, localized }) = |
|
348 |
mk_class_data (f ((locale, consts, v, inst, intro), (defs, localized))) |
|
349 |
fun merge_class_data _ (ClassData { locale = locale, consts = consts, v = v, inst = inst, |
|
350 |
intro = intro, defs = defs1, localized = localized1 }, |
|
351 |
ClassData { locale = _, consts = _, v = _, inst = _, intro = _, |
|
352 |
defs = defs2, localized = localized2 }) = |
|
353 |
mk_class_data ((locale, consts, v, inst, intro), |
|
354 |
(Thm.merge_thms (defs1, defs2), AList.merge (op =) (K true) (localized1, localized2))); |
|
24218 | 355 |
|
356 |
fun merge_pair f1 f2 ((x1, y1), (x2, y2)) = (f1 (x1, x2), f2 (y1, y2)); |
|
357 |
||
358 |
structure ClassData = TheoryDataFun |
|
359 |
( |
|
24657 | 360 |
type T = class_data Graph.T * class Symtab.table |
361 |
(*locale name ~> class name*); |
|
24218 | 362 |
val empty = (Graph.empty, Symtab.empty); |
363 |
val copy = I; |
|
364 |
val extend = I; |
|
24657 | 365 |
fun merge _ = merge_pair (Graph.join merge_class_data) (Symtab.merge (K true)); |
24218 | 366 |
); |
367 |
||
368 |
||
369 |
(* queries *) |
|
370 |
||
24657 | 371 |
val lookup_class_data = Option.map rep_class_data oo try o Graph.get_node |
24589 | 372 |
o fst o ClassData.get; |
24218 | 373 |
fun class_of_locale thy = Symtab.lookup ((snd o ClassData.get) thy); |
374 |
||
24589 | 375 |
fun the_class_data thy class = case lookup_class_data thy class |
376 |
of NONE => error ("undeclared class " ^ quote class) |
|
377 |
| SOME data => data; |
|
24218 | 378 |
|
379 |
val ancestry = Graph.all_succs o fst o ClassData.get; |
|
380 |
||
24435 | 381 |
fun params_of_sort thy = |
24218 | 382 |
let |
383 |
fun params class = |
|
384 |
let |
|
385 |
val const_typs = (#params o AxClass.get_definition thy) class; |
|
24657 | 386 |
val const_names = (#consts o the_class_data thy) class; |
24218 | 387 |
in |
388 |
(map o apsnd) (fn c => (c, (the o AList.lookup (op =) const_typs) c)) const_names |
|
389 |
end; |
|
390 |
in maps params o ancestry thy end; |
|
391 |
||
24657 | 392 |
fun these_defs thy = maps (these o Option.map #defs o lookup_class_data thy) o ancestry thy; |
24218 | 393 |
|
394 |
fun these_intros thy = |
|
24657 | 395 |
Graph.fold (fn (_, (data, _)) => insert Thm.eq_thm ((#intro o rep_class_data) data)) |
24218 | 396 |
((fst o ClassData.get) thy) []; |
397 |
||
24657 | 398 |
fun these_localized thy class = |
399 |
maps (#localized o the_class_data thy) (ancestry thy [class]); |
|
400 |
||
401 |
fun local_param thy = AList.lookup (op =) o these_localized thy; |
|
402 |
||
24218 | 403 |
fun print_classes thy = |
404 |
let |
|
405 |
val algebra = Sign.classes_of thy; |
|
406 |
val arities = |
|
407 |
Symtab.empty |
|
408 |
|> Symtab.fold (fn (tyco, arities) => fold (fn (class, _) => |
|
409 |
Symtab.map_default (class, []) (insert (op =) tyco)) arities) |
|
410 |
((#arities o Sorts.rep_algebra) algebra); |
|
411 |
val the_arities = these o Symtab.lookup arities; |
|
412 |
fun mk_arity class tyco = |
|
413 |
let |
|
414 |
val Ss = Sorts.mg_domain algebra tyco [class]; |
|
415 |
in Sign.pretty_arity thy (tyco, Ss, [class]) end; |
|
416 |
fun mk_param (c, ty) = Pretty.str (Sign.extern_const thy c ^ " :: " |
|
417 |
^ setmp show_sorts false (Sign.string_of_typ thy o Type.strip_sorts) ty); |
|
418 |
fun mk_entry class = (Pretty.block o Pretty.fbreaks o map_filter I) [ |
|
419 |
(SOME o Pretty.str) ("class " ^ class ^ ":"), |
|
420 |
(SOME o Pretty.block) [Pretty.str "supersort: ", |
|
421 |
(Sign.pretty_sort thy o Sign.certify_sort thy o Sign.super_classes thy) class], |
|
24657 | 422 |
Option.map (Pretty.str o prefix "locale: " o #locale) (lookup_class_data thy class), |
24218 | 423 |
((fn [] => NONE | ps => (SOME o Pretty.block o Pretty.fbreaks) (Pretty.str "parameters:" :: ps)) o map mk_param |
424 |
o these o Option.map #params o try (AxClass.get_definition thy)) class, |
|
425 |
(SOME o Pretty.block o Pretty.breaks) [ |
|
426 |
Pretty.str "instances:", |
|
427 |
Pretty.list "" "" (map (mk_arity class) (the_arities class)) |
|
428 |
] |
|
429 |
] |
|
430 |
in |
|
24589 | 431 |
(Pretty.writeln o Pretty.chunks o separate (Pretty.str "") |
432 |
o map mk_entry o Sorts.all_classes) algebra |
|
24218 | 433 |
end; |
434 |
||
435 |
||
436 |
(* updaters *) |
|
437 |
||
24589 | 438 |
fun add_class_data ((class, superclasses), (locale, consts, v, inst, intro)) = |
24218 | 439 |
ClassData.map (fn (gr, tab) => ( |
440 |
gr |
|
24657 | 441 |
|> Graph.new_node (class, mk_class_data ((locale, (map o apfst) fst consts, v, inst, intro), |
442 |
([], map (apsnd (rpair (class, 0) o Free) o swap) consts))) |
|
24218 | 443 |
|> fold (curry Graph.add_edge class) superclasses, |
444 |
tab |
|
445 |
|> Symtab.update (locale, class) |
|
446 |
)); |
|
447 |
||
24657 | 448 |
fun add_class_const_def (class, (entry, def)) = |
449 |
(ClassData.map o apfst o Graph.map_node class o map_class_data o apsnd) |
|
450 |
(fn (defs, localized) => (def :: defs, (apsnd o apsnd) (pair class) entry :: localized)); |
|
24218 | 451 |
|
24589 | 452 |
|
453 |
(** rule calculation, tactics and methods **) |
|
454 |
||
455 |
fun class_intro thy locale class sups = |
|
456 |
let |
|
457 |
fun class_elim class = |
|
458 |
case (map Drule.unconstrainTs o #axioms o AxClass.get_definition thy) class |
|
459 |
of [thm] => SOME thm |
|
460 |
| [] => NONE; |
|
461 |
val pred_intro = case Locale.intros thy locale |
|
462 |
of ([ax_intro], [intro]) => intro |> OF_LAST ax_intro |> SOME |
|
463 |
| ([intro], []) => SOME intro |
|
464 |
| ([], [intro]) => SOME intro |
|
465 |
| _ => NONE; |
|
466 |
val pred_intro' = pred_intro |
|
467 |
|> Option.map (fn intro => intro OF map_filter class_elim sups); |
|
468 |
val class_intro = (#intro o AxClass.get_definition thy) class; |
|
469 |
val raw_intro = case pred_intro' |
|
470 |
of SOME pred_intro => class_intro |> OF_LAST pred_intro |
|
471 |
| NONE => class_intro; |
|
472 |
val sort = Sign.super_classes thy class; |
|
473 |
val typ = TVar ((AxClass.param_tyvarname, 0), sort); |
|
474 |
val defs = these_defs thy sups; |
|
475 |
in |
|
476 |
raw_intro |
|
477 |
|> Drule.instantiate' [SOME (Thm.ctyp_of thy typ)] [] |
|
478 |
|> strip_all_ofclass thy sort |
|
479 |
|> Thm.strip_shyps |
|
480 |
|> MetaSimplifier.rewrite_rule defs |
|
481 |
|> Drule.unconstrainTs |
|
482 |
end; |
|
483 |
||
484 |
fun class_interpretation class facts defs thy = |
|
485 |
let |
|
24657 | 486 |
val { locale, inst, ... } = the_class_data thy class; |
24589 | 487 |
val tac = (ALLGOALS o ProofContext.fact_tac) facts; |
488 |
val prfx = Logic.const_of_class (NameSpace.base class); |
|
489 |
in |
|
490 |
prove_interpretation tac ((false, prfx), []) (Locale.Locale locale) |
|
491 |
(inst, defs) thy |
|
492 |
end; |
|
493 |
||
494 |
fun interpretation_in_rule thy (class1, class2) = |
|
495 |
let |
|
496 |
fun mk_axioms class = |
|
497 |
let |
|
24657 | 498 |
val { locale, inst = (_, insttab), ... } = the_class_data thy class; |
24589 | 499 |
in |
500 |
Locale.global_asms_of thy locale |
|
501 |
|> maps snd |
|
502 |
|> (map o map_aterms) (fn Free (s, _) => (the o Symtab.lookup insttab) s | t => t) |
|
503 |
|> (map o map_types o map_atyps) (fn TFree _ => TFree ("'a", [class1]) | T => T) |
|
504 |
|> map (ObjectLogic.ensure_propT thy) |
|
505 |
end; |
|
506 |
val (prems, concls) = pairself mk_axioms (class1, class2); |
|
507 |
in |
|
508 |
Goal.prove_global thy [] prems (Logic.mk_conjunction_list concls) |
|
509 |
(Locale.intro_locales_tac true (ProofContext.init thy)) |
|
510 |
end; |
|
24218 | 511 |
|
512 |
fun intro_classes_tac facts st = |
|
513 |
let |
|
514 |
val thy = Thm.theory_of_thm st; |
|
515 |
val classes = Sign.all_classes thy; |
|
516 |
val class_trivs = map (Thm.class_triv thy) classes; |
|
517 |
val class_intros = these_intros thy; |
|
518 |
fun add_axclass_intro class = |
|
519 |
case try (AxClass.get_definition thy) class of SOME {intro, ...} => cons intro | _ => I; |
|
520 |
val axclass_intros = fold add_axclass_intro classes []; |
|
521 |
in |
|
522 |
st |
|
523 |
|> ((ALLGOALS (Method.insert_tac facts THEN' |
|
524 |
REPEAT_ALL_NEW (resolve_tac (class_trivs @ class_intros @ axclass_intros)))) |
|
525 |
THEN Tactic.distinct_subgoals_tac) |
|
526 |
end; |
|
527 |
||
528 |
fun default_intro_classes_tac [] = intro_classes_tac [] |
|
529 |
| default_intro_classes_tac _ = Tactical.no_tac; (*no error message!*) |
|
530 |
||
531 |
fun default_tac rules ctxt facts = |
|
532 |
HEADGOAL (Method.some_rule_tac rules ctxt facts) ORELSE |
|
533 |
default_intro_classes_tac facts; |
|
534 |
||
535 |
val _ = Context.add_setup (Method.add_methods |
|
536 |
[("intro_classes", Method.no_args (Method.METHOD intro_classes_tac), |
|
537 |
"back-chain introduction rules of classes"), |
|
538 |
("default", Method.thms_ctxt_args (Method.METHOD oo default_tac), |
|
539 |
"apply some intro/elim rule")]); |
|
540 |
||
541 |
||
24589 | 542 |
(** classes and class target **) |
24218 | 543 |
|
24589 | 544 |
(* class definition *) |
24218 | 545 |
|
546 |
local |
|
547 |
||
24707 | 548 |
fun read_param thy raw_t = (* FIXME ProofContext.read_const (!?) *) |
24218 | 549 |
let |
24707 | 550 |
val t = Syntax.read_term_global thy raw_t |
24218 | 551 |
in case try dest_Const t |
552 |
of SOME (c, _) => c |
|
553 |
| NONE => error ("Not a constant: " ^ Sign.string_of_term thy t) |
|
554 |
end; |
|
555 |
||
556 |
fun gen_class add_locale prep_class prep_param bname |
|
557 |
raw_supclasses raw_elems raw_other_consts thy = |
|
558 |
let |
|
24589 | 559 |
fun mk_instT class = Symtab.empty |
560 |
|> Symtab.update (AxClass.param_tyvarname, TFree (AxClass.param_tyvarname, [class])); |
|
561 |
fun mk_inst class param_names cs = |
|
562 |
Symtab.empty |
|
563 |
|> fold2 (fn v => fn (c, ty) => Symtab.update (v, Const |
|
564 |
(c, Term.map_type_tfree (fn (v, _) => TFree (v, [class])) ty))) param_names cs; |
|
24218 | 565 |
(*FIXME need proper concept for reading locale statements*) |
566 |
fun subst_classtyvar (_, _) = |
|
567 |
TFree (AxClass.param_tyvarname, []) |
|
568 |
| subst_classtyvar (v, sort) = |
|
24589 | 569 |
error ("Sort constraint illegal in type class, for type variable " |
570 |
^ v ^ "::" ^ Sign.string_of_sort thy sort); |
|
24218 | 571 |
(*val subst_classtyvars = Element.map_ctxt {name = I, var = I, term = I, |
572 |
typ = Term.map_type_tfree subst_classtyvar, fact = I, attrib = I};*) |
|
573 |
val other_consts = map (prep_param thy) raw_other_consts; |
|
574 |
val (elems, includes) = fold_rev (fn Locale.Elem e => apfst (cons e) |
|
575 |
| Locale.Expr i => apsnd (cons i)) raw_elems ([], []); |
|
576 |
val supclasses = map (prep_class thy) raw_supclasses; |
|
577 |
val sups = filter (is_some o lookup_class_data thy) supclasses |
|
578 |
|> Sign.certify_sort thy; |
|
579 |
val supsort = Sign.certify_sort thy supclasses; |
|
24657 | 580 |
val suplocales = map (Locale.Locale o #locale o the_class_data thy) sups; |
24218 | 581 |
val supexpr = Locale.Merge (suplocales @ includes); |
582 |
val supparams = (map fst o Locale.parameters_of_expr thy) |
|
583 |
(Locale.Merge suplocales); |
|
24435 | 584 |
val supconsts = AList.make (the o AList.lookup (op =) (params_of_sort thy sups)) |
24218 | 585 |
(map fst supparams); |
586 |
(*val elems_constrains = map |
|
587 |
(Element.Constrains o apsnd (Term.map_type_tfree subst_classtyvar)) supparams;*) |
|
588 |
fun mk_tyvar (_, sort) = TFree (AxClass.param_tyvarname, |
|
589 |
if Sign.subsort thy (supsort, sort) then sort else error |
|
590 |
("Sort " ^ Sign.string_of_sort thy sort |
|
591 |
^ " is less general than permitted least general sort " |
|
592 |
^ Sign.string_of_sort thy supsort)); |
|
593 |
fun extract_params thy name_locale = |
|
594 |
let |
|
595 |
val params = Locale.parameters_of thy name_locale; |
|
596 |
val v = case (maps typ_tfrees o map (snd o fst)) params |
|
24657 | 597 |
of (v, _) :: _ => v |
598 |
| [] => AxClass.param_tyvarname; |
|
24218 | 599 |
in |
24657 | 600 |
(v, (map fst params, params |
24218 | 601 |
|> (map o apfst o apsnd o Term.map_type_tfree) mk_tyvar |
602 |
|> (map o apsnd) (fork_mixfix true NONE #> fst) |
|
603 |
|> chop (length supconsts) |
|
604 |
|> snd)) |
|
605 |
end; |
|
606 |
fun extract_assumes name_locale params thy cs = |
|
607 |
let |
|
608 |
val consts = supconsts @ (map (fst o fst) params ~~ cs); |
|
609 |
fun subst (Free (c, ty)) = |
|
610 |
Const ((fst o the o AList.lookup (op =) consts) c, ty) |
|
611 |
| subst t = t; |
|
612 |
fun prep_asm ((name, atts), ts) = |
|
613 |
((NameSpace.base name, map (Attrib.attribute thy) atts), |
|
24589 | 614 |
(map o map_aterms) subst ts); |
24218 | 615 |
in |
616 |
Locale.global_asms_of thy name_locale |
|
617 |
|> map prep_asm |
|
618 |
end; |
|
619 |
fun note_intro name_axclass class_intro = |
|
620 |
PureThy.note_thmss_qualified "" ((Logic.const_of_class o NameSpace.base) name_axclass) |
|
621 |
[(("intro", []), [([class_intro], [])])] |
|
622 |
#> snd |
|
623 |
in |
|
624 |
thy |
|
625 |
|> add_locale (SOME "") bname supexpr ((*elems_constrains @*) elems) |
|
626 |
|-> (fn name_locale => ProofContext.theory_result ( |
|
627 |
`(fn thy => extract_params thy name_locale) |
|
24657 | 628 |
#-> (fn (v, (globals, params)) => |
24589 | 629 |
AxClass.define_class_params (bname, supsort) params |
630 |
(extract_assumes name_locale params) other_consts |
|
631 |
#-> (fn (name_axclass, (consts, axioms)) => |
|
24218 | 632 |
`(fn thy => class_intro thy name_locale name_axclass sups) |
633 |
#-> (fn class_intro => |
|
634 |
add_class_data ((name_axclass, sups), |
|
24657 | 635 |
(name_locale, map fst params ~~ map fst consts, v, |
636 |
(mk_instT name_axclass, mk_inst name_axclass (map fst globals) |
|
24589 | 637 |
(map snd supconsts @ consts)), class_intro)) |
24218 | 638 |
#> note_intro name_axclass class_intro |
24589 | 639 |
#> class_interpretation name_axclass axioms [] |
24218 | 640 |
#> pair name_axclass |
641 |
))))) |
|
642 |
end; |
|
643 |
||
644 |
in |
|
645 |
||
646 |
val class_cmd = gen_class Locale.add_locale Sign.intern_class read_param; |
|
647 |
val class = gen_class Locale.add_locale_i Sign.certify_class (K I); |
|
648 |
||
649 |
end; (*local*) |
|
650 |
||
651 |
||
24657 | 652 |
(* class target context *) |
653 |
||
654 |
fun remove_constraints class thy = |
|
655 |
thy |> fold_map (get_remove_contraint o fst) (these_localized thy class); |
|
656 |
||
657 |
||
24589 | 658 |
(* definition in class target *) |
24218 | 659 |
|
660 |
fun export_fixes thy class = |
|
661 |
let |
|
24435 | 662 |
val consts = params_of_sort thy [class]; |
24218 | 663 |
fun subst_aterm (t as Free (v, ty)) = (case AList.lookup (op =) consts v |
664 |
of SOME (c, _) => Const (c, ty) |
|
665 |
| NONE => t) |
|
666 |
| subst_aterm t = t; |
|
24657 | 667 |
in Term.map_aterms subst_aterm end; |
24218 | 668 |
|
669 |
fun add_const_in_class class ((c, rhs), syn) thy = |
|
670 |
let |
|
671 |
val prfx = (Logic.const_of_class o NameSpace.base) class; |
|
24657 | 672 |
fun mk_name c = |
24218 | 673 |
let |
674 |
val n1 = Sign.full_name thy c; |
|
675 |
val n2 = NameSpace.qualifier n1; |
|
676 |
val n3 = NameSpace.base n1; |
|
24657 | 677 |
in NameSpace.implode [n2, prfx, n3] end; |
678 |
val v = (#v o the_class_data thy) class; |
|
679 |
val constrain_sort = curry (Sorts.inter_sort (Sign.classes_of thy)) [class]; |
|
680 |
val subst_typ = Term.map_type_tfree (fn var as (w, sort) => |
|
681 |
if w = v then TFree (w, constrain_sort sort) else TFree var); |
|
24218 | 682 |
val rhs' = export_fixes thy class rhs; |
683 |
val ty' = Term.fastype_of rhs'; |
|
24657 | 684 |
val ty'' = subst_typ ty'; |
685 |
val c' = mk_name c; |
|
686 |
val def = (c, Logic.mk_equals (Const (c', ty'), rhs')); |
|
24218 | 687 |
val (syn', _) = fork_mixfix true NONE syn; |
24657 | 688 |
fun interpret def thy = |
24218 | 689 |
let |
24589 | 690 |
val def' = symmetric def; |
24218 | 691 |
val def_eq = Thm.prop_of def'; |
24657 | 692 |
val typargs = Sign.const_typargs thy (c', fastype_of rhs); |
693 |
val typidx = find_index (fn TFree (w, _) => v = w | _ => false) typargs; |
|
24218 | 694 |
in |
24657 | 695 |
thy |
696 |
|> class_interpretation class [def'] [def_eq] |
|
697 |
|> add_class_const_def (class, ((c', (rhs, typidx)), def')) |
|
24218 | 698 |
end; |
699 |
in |
|
700 |
thy |
|
701 |
|> Sign.add_path prfx |
|
702 |
|> Sign.add_consts_authentic [(c, ty', syn')] |
|
703 |
|> Sign.parent_path |
|
704 |
|> Sign.sticky_prefix prfx |
|
705 |
|> PureThy.add_defs_i false [(def, [])] |
|
706 |
|-> (fn [def] => interpret def) |
|
24657 | 707 |
|> Sign.add_const_constraint_i (c', SOME ty'') |
24218 | 708 |
|> Sign.restore_naming thy |
709 |
end; |
|
710 |
||
24589 | 711 |
|
712 |
(* interpretation in class target *) |
|
713 |
||
714 |
local |
|
715 |
||
716 |
fun gen_interpretation_in_class prep_class do_proof (raw_class, raw_superclass) theory = |
|
717 |
let |
|
718 |
val class = prep_class theory raw_class; |
|
719 |
val superclass = prep_class theory raw_superclass; |
|
24657 | 720 |
val loc_name = (#locale o the_class_data theory) class; |
721 |
val loc_expr = (Locale.Locale o #locale o the_class_data theory) superclass; |
|
24589 | 722 |
fun prove_classrel (class, superclass) thy = |
723 |
let |
|
724 |
val classes = (Graph.all_succs o #classes o Sorts.rep_algebra |
|
725 |
o Sign.classes_of) thy [superclass] |
|
726 |
|> filter_out (fn class' => Sign.subsort thy ([class], [class'])); |
|
727 |
fun instance_subclass (class1, class2) thy = |
|
728 |
let |
|
729 |
val interp = interpretation_in_rule thy (class1, class2); |
|
730 |
val ax = #axioms (AxClass.get_definition thy class1); |
|
731 |
val intro = #intro (AxClass.get_definition thy class2) |
|
732 |
|> Drule.instantiate' [SOME (Thm.ctyp_of thy |
|
733 |
(TVar ((AxClass.param_tyvarname, 0), [class1])))] []; |
|
734 |
val thm = |
|
735 |
intro |
|
736 |
|> OF_LAST (interp OF ax) |
|
737 |
|> strip_all_ofclass thy (Sign.super_classes thy class2); |
|
738 |
in |
|
739 |
thy |> AxClass.add_classrel thm |
|
740 |
end; |
|
741 |
in |
|
742 |
thy |> fold_rev (curry instance_subclass class) classes |
|
743 |
end; |
|
744 |
in |
|
745 |
theory |
|
746 |
|> do_proof (prove_classrel (class, superclass)) (loc_name, loc_expr) |
|
747 |
end; |
|
748 |
||
749 |
in |
|
750 |
||
751 |
val interpretation_in_class = gen_interpretation_in_class Sign.certify_class |
|
752 |
(Locale.interpretation_in_locale o ProofContext.theory); |
|
753 |
val interpretation_in_class_cmd = gen_interpretation_in_class Sign.read_class |
|
754 |
(Locale.interpretation_in_locale o ProofContext.theory); |
|
755 |
val prove_interpretation_in_class = gen_interpretation_in_class Sign.certify_class |
|
756 |
o prove_interpretation_in; |
|
757 |
||
758 |
end; (*local*) |
|
759 |
||
24657 | 760 |
(*experimental*) |
761 |
fun class_term_check thy class = |
|
762 |
let |
|
763 |
val algebra = Sign.classes_of thy; |
|
764 |
val { v, ... } = the_class_data thy class; |
|
765 |
fun add_constrain_classtyp sort' (ty as TFree (v, _)) = |
|
766 |
AList.map_default (op =) (v, []) (curry (Sorts.inter_sort algebra) sort') |
|
767 |
| add_constrain_classtyp sort' (Type (tyco, tys)) = case Sorts.mg_domain algebra tyco sort' |
|
768 |
of sorts => fold2 add_constrain_classtyp sorts tys; |
|
769 |
fun class_arg c idx ty = |
|
770 |
let |
|
771 |
val typargs = Sign.const_typargs thy (c, ty); |
|
772 |
fun classtyp (t as TFree (w, _)) = if w = v then NONE else SOME t |
|
773 |
| classtyp t = SOME t; |
|
774 |
in classtyp (nth typargs idx) end; |
|
775 |
fun add_inst (c, ty) (terminsts, typinsts) = case local_param thy class c |
|
776 |
of NONE => (terminsts, typinsts) |
|
777 |
| SOME (t, (class', idx)) => (case class_arg c idx ty |
|
778 |
of NONE => (((c, ty), t) :: terminsts, typinsts) |
|
779 |
| SOME ty => (terminsts, add_constrain_classtyp [class'] ty typinsts)); |
|
780 |
in pair o (fn ts => let |
|
781 |
val cs = (fold o fold_aterms) (fn Const c_ty => insert (op =) c_ty | _ => I) ts []; |
|
782 |
val (terminsts, typinsts) = fold add_inst cs ([], []); |
|
783 |
in |
|
784 |
ts |
|
785 |
|> (map o map_aterms) (fn t as Const c_ty => the_default t (AList.lookup (op =) terminsts c_ty) |
|
786 |
| t => t) |
|
787 |
|> (map o map_types o map_atyps) (fn t as TFree (v, sort) => |
|
788 |
case AList.lookup (op =) typinsts v |
|
789 |
of SOME sort' => TFree (v, Sorts.inter_sort algebra (sort, sort')) |
|
790 |
| NONE => t) |
|
791 |
end) end; |
|
792 |
||
793 |
val init_ref = ref (K (pair I) : class -> Proof.context -> (theory -> theory) * Proof.context); |
|
794 |
fun init class = ! init_ref class; |
|
795 |
||
796 |
fun init_default class ctxt = |
|
797 |
let |
|
798 |
val thy = ProofContext.theory_of ctxt; |
|
799 |
val term_check = class_term_check thy class; |
|
800 |
in |
|
801 |
ctxt |
|
802 |
(*|> ProofContext.theory_result (remove_constraints class)*) |
|
803 |
|> Context.proof_map (Syntax.add_term_check term_check) |
|
804 |
(*|>> fold (fn (c, ty) => Sign.add_const_constraint_i (c, SOME ty))*) |
|
805 |
|> pair I |
|
806 |
end; |
|
807 |
||
24218 | 808 |
end; |