author | haftmann |
Mon, 02 Oct 2006 23:00:51 +0200 | |
changeset 20835 | 27d049062b56 |
parent 20703 | f3f2b1091ea0 |
child 20843 | a5343075bdc5 |
permissions | -rw-r--r-- |
18168 | 1 |
(* Title: Pure/Tools/class_package.ML |
2 |
ID: $Id$ |
|
3 |
Author: Florian Haftmann, TU Muenchen |
|
4 |
||
19468 | 5 |
Type classes derived from primitive axclasses and locales. |
18168 | 6 |
*) |
7 |
||
8 |
signature CLASS_PACKAGE = |
|
9 |
sig |
|
20428 | 10 |
val class: bstring -> class list -> Element.context Locale.element list -> theory |
20289 | 11 |
-> Proof.context * theory |
20428 | 12 |
val class_i: bstring -> class list -> Element.context_i Locale.element list -> theory |
20289 | 13 |
-> Proof.context * theory |
20601 | 14 |
val instance_arity: ((bstring * string list) * string) list |
19136 | 15 |
-> bstring * Attrib.src list -> ((bstring * Attrib.src list) * string) list |
18575 | 16 |
-> theory -> Proof.state |
20601 | 17 |
val instance_arity_i: ((bstring * sort list) * sort) list |
19136 | 18 |
-> bstring * attribute list -> ((bstring * attribute list) * term) list |
18575 | 19 |
-> theory -> Proof.state |
20182
79c9ff40d760
tactic for prove_instance_arity now gets definition theorems as arguments
haftmann
parents:
20175
diff
changeset
|
20 |
val prove_instance_arity: (thm list -> tactic) -> ((string * sort list) * sort) list |
19136 | 21 |
-> bstring * attribute list -> ((bstring * attribute list) * term) list |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
22 |
-> theory -> theory |
19150 | 23 |
val instance_sort: string * string -> theory -> Proof.state |
24 |
val instance_sort_i: class * sort -> theory -> Proof.state |
|
25 |
val prove_instance_sort: tactic -> class * sort -> theory -> theory |
|
18168 | 26 |
|
18755 | 27 |
val certify_class: theory -> class -> class |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
28 |
val certify_sort: theory -> sort -> sort |
20385 | 29 |
val read_class: theory -> xstring -> class |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
30 |
val read_sort: theory -> string -> sort |
20465 | 31 |
val operational_algebra: theory -> (sort -> sort) * Sorts.algebra |
18702 | 32 |
val the_consts_sign: theory -> class -> string * (string * typ) list |
33 |
val the_inst_sign: theory -> class * string -> (string * sort) list * (string * typ) list |
|
20175 | 34 |
val assume_arities_of_sort: theory -> ((string * sort list) * sort) list -> typ * sort -> bool |
35 |
val assume_arities_thy: theory -> ((string * sort list) * sort) list -> (theory -> 'a) -> 'a |
|
36 |
(*'a must not keep any reference to theory*) |
|
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
37 |
|
18702 | 38 |
val print_classes: theory -> unit |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
39 |
val intro_classes_tac: thm list -> tactic |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
40 |
val default_intro_classes_tac: thm list -> tactic |
18168 | 41 |
end; |
42 |
||
20182
79c9ff40d760
tactic for prove_instance_arity now gets definition theorems as arguments
haftmann
parents:
20175
diff
changeset
|
43 |
structure ClassPackage : CLASS_PACKAGE = |
18168 | 44 |
struct |
45 |
||
46 |
||
20455 | 47 |
(** theory data **) |
18168 | 48 |
|
19456 | 49 |
datatype class_data = ClassData of { |
18575 | 50 |
name_locale: string, |
51 |
name_axclass: string, |
|
52 |
var: string, |
|
19966 | 53 |
consts: (string * (string * typ)) list, |
20385 | 54 |
(*locale parameter ~> toplevel theory constant*) |
19966 | 55 |
propnames: string list |
56 |
} * thm list Symtab.table; |
|
18168 | 57 |
|
19456 | 58 |
fun rep_classdata (ClassData c) = c; |
59 |
||
18575 | 60 |
structure ClassData = TheoryDataFun ( |
18168 | 61 |
struct |
62 |
val name = "Pure/classes"; |
|
19966 | 63 |
type T = class_data Symtab.table; |
64 |
val empty = Symtab.empty; |
|
18168 | 65 |
val copy = I; |
66 |
val extend = I; |
|
19966 | 67 |
fun merge _ = Symtab.join (fn _ => fn (ClassData (classd, instd1), ClassData (_, instd2)) => |
68 |
(ClassData (classd, Symtab.merge (K true) (instd1, instd2)))); |
|
69 |
fun print thy data = |
|
18575 | 70 |
let |
19966 | 71 |
fun pretty_class (name, ClassData ({name_locale, name_axclass, var, consts, ...}, _)) = |
18575 | 72 |
(Pretty.block o Pretty.fbreaks) [ |
73 |
Pretty.str ("class " ^ name ^ ":"), |
|
74 |
Pretty.str ("locale: " ^ name_locale), |
|
75 |
Pretty.str ("axclass: " ^ name_axclass), |
|
76 |
Pretty.str ("class variable: " ^ var), |
|
77 |
(Pretty.block o Pretty.fbreaks) ( |
|
78 |
Pretty.str "constants: " |
|
19280
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
79 |
:: map (fn (_, (c, ty)) => Pretty.str (c ^ " :: " ^ Sign.string_of_typ thy ty)) consts |
18575 | 80 |
) |
81 |
] |
|
82 |
in |
|
19966 | 83 |
(Pretty.writeln o Pretty.chunks o map pretty_class o Symtab.dest) data |
18575 | 84 |
end; |
18168 | 85 |
end |
86 |
); |
|
87 |
||
18708 | 88 |
val _ = Context.add_setup ClassData.init; |
18575 | 89 |
val print_classes = ClassData.print; |
90 |
||
19038 | 91 |
|
92 |
(* queries *) |
|
93 |
||
19966 | 94 |
val lookup_class_data = Option.map rep_classdata oo Symtab.lookup o ClassData.get; |
18168 | 95 |
|
19038 | 96 |
fun the_class_data thy class = |
18168 | 97 |
case lookup_class_data thy class |
20455 | 98 |
of NONE => error ("undeclared constructive class " ^ quote class) |
18168 | 99 |
| SOME data => data; |
100 |
||
19280
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
101 |
fun is_operational_class thy cls = |
18702 | 102 |
lookup_class_data thy cls |
19966 | 103 |
|> Option.map (not o null o #consts o fst) |
18702 | 104 |
|> the_default false; |
18168 | 105 |
|
20455 | 106 |
fun the_ancestry thy classes = |
18360 | 107 |
let |
20455 | 108 |
fun proj_class class = |
19280
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
109 |
if is_operational_class thy class |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
110 |
then [class] |
20455 | 111 |
else (Sign.certify_sort thy o maps proj_class o Sign.super_classes thy) class; |
19038 | 112 |
fun ancestry class anc = |
113 |
anc |
|
20385 | 114 |
|> insert (op =) class |
20455 | 115 |
|> fold ancestry ((maps proj_class o Sign.super_classes thy) class); |
19038 | 116 |
in fold ancestry classes [] end; |
117 |
||
20455 | 118 |
val the_parm_map = #consts o fst oo the_class_data; |
119 |
||
120 |
val the_propnames = #propnames o fst oo the_class_data; |
|
121 |
||
19038 | 122 |
fun subst_clsvar v ty_subst = |
123 |
map_type_tfree (fn u as (w, _) => |
|
124 |
if w = v then ty_subst else TFree u); |
|
125 |
||
18168 | 126 |
|
19038 | 127 |
(* updaters *) |
128 |
||
19966 | 129 |
fun add_class_data (class, (name_locale, name_axclass, var, consts, propnames)) = |
19957 | 130 |
ClassData.map ( |
19966 | 131 |
Symtab.update_new (class, ClassData ({ |
19957 | 132 |
name_locale = name_locale, |
133 |
name_axclass = name_axclass, |
|
134 |
var = var, |
|
19966 | 135 |
consts = consts, |
136 |
propnames = propnames}, Symtab.empty)) |
|
19957 | 137 |
); |
19038 | 138 |
|
19966 | 139 |
fun add_inst_def ((class, tyco), thm) = |
140 |
ClassData.map ( |
|
141 |
Symtab.map_entry class (fn ClassData (classd, instd) => |
|
142 |
ClassData (classd, Symtab.insert_list eq_thm (tyco, thm) instd)) |
|
143 |
); |
|
19038 | 144 |
|
20106 | 145 |
|
20385 | 146 |
(* certification and reading *) |
19038 | 147 |
|
148 |
fun certify_class thy class = |
|
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
149 |
(fn class => (the_class_data thy class; class)) (Sign.certify_class thy class); |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
150 |
|
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
151 |
fun certify_sort thy sort = |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
152 |
map (fn class => (the_class_data thy class; class)) (Sign.certify_sort thy sort); |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
153 |
|
20385 | 154 |
fun read_class thy = |
19957 | 155 |
certify_class thy o Sign.intern_class thy; |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
156 |
|
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
157 |
fun read_sort thy = |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
158 |
certify_sort thy o Sign.read_sort thy; |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
159 |
|
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
160 |
|
20455 | 161 |
|
162 |
(** contexts with arity assumptions **) |
|
20175 | 163 |
|
164 |
fun assume_arities_of_sort thy arities ty_sort = |
|
165 |
let |
|
166 |
val pp = Sign.pp thy; |
|
167 |
val algebra = Sign.classes_of thy |
|
168 |
|> fold (fn ((tyco, asorts), sort) => |
|
169 |
Sorts.add_arities pp (tyco, map (fn class => (class, asorts)) sort)) arities; |
|
170 |
in Sorts.of_sort algebra ty_sort end; |
|
171 |
||
172 |
fun assume_arities_thy thy arities f = |
|
173 |
let |
|
174 |
val thy_read = (fold (fn ((tyco, asorts), sort) |
|
175 |
=> Sign.primitive_arity (tyco, asorts, sort)) arities o Theory.copy) thy |
|
176 |
in f thy_read end; |
|
177 |
||
20455 | 178 |
|
179 |
||
180 |
(** tactics and methods **) |
|
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
181 |
|
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
182 |
fun intro_classes_tac facts st = |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
183 |
(ALLGOALS (Method.insert_tac facts THEN' |
19928 | 184 |
REPEAT_ALL_NEW (resolve_tac (AxClass.class_intros (Thm.theory_of_thm st)))) |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
185 |
THEN Tactic.distinct_subgoals_tac) st; |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
186 |
|
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
187 |
fun default_intro_classes_tac [] = intro_classes_tac [] |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
188 |
| default_intro_classes_tac _ = Tactical.no_tac; (*no error message!*) |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
189 |
|
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
190 |
fun default_tac rules ctxt facts = |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
191 |
HEADGOAL (Method.some_rule_tac rules ctxt facts) ORELSE |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
192 |
default_intro_classes_tac facts; |
19038 | 193 |
|
19468 | 194 |
val _ = Context.add_setup (Method.add_methods |
195 |
[("intro_classes", Method.no_args (Method.METHOD intro_classes_tac), |
|
196 |
"back-chain introduction rules of classes"), |
|
197 |
("default", Method.thms_ctxt_args (Method.METHOD oo default_tac), |
|
198 |
"apply some intro/elim rule")]); |
|
199 |
||
19038 | 200 |
|
20455 | 201 |
|
202 |
(** axclass instances **) |
|
19246 | 203 |
|
204 |
local |
|
205 |
||
20175 | 206 |
fun gen_instance mk_prop add_thm after_qed insts thy = |
20368
2461a0485623
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
wenzelm
parents:
20353
diff
changeset
|
207 |
let |
2461a0485623
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
wenzelm
parents:
20353
diff
changeset
|
208 |
fun after_qed' results = |
2461a0485623
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
wenzelm
parents:
20353
diff
changeset
|
209 |
ProofContext.theory ((fold o fold) add_thm results #> after_qed); |
2461a0485623
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
wenzelm
parents:
20353
diff
changeset
|
210 |
in |
2461a0485623
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
wenzelm
parents:
20353
diff
changeset
|
211 |
thy |
2461a0485623
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
wenzelm
parents:
20353
diff
changeset
|
212 |
|> ProofContext.init |
2461a0485623
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
wenzelm
parents:
20353
diff
changeset
|
213 |
|> Proof.theorem_i PureThy.internalK NONE after_qed' NONE ("", []) |
2461a0485623
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
wenzelm
parents:
20353
diff
changeset
|
214 |
((map (fn t => (("", []), [(t, [])])) o maps (mk_prop thy)) insts) |
2461a0485623
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
wenzelm
parents:
20353
diff
changeset
|
215 |
end; |
19246 | 216 |
|
217 |
in |
|
218 |
||
219 |
val axclass_instance_arity = |
|
19412 | 220 |
gen_instance (Logic.mk_arities oo Sign.read_arity) AxClass.add_arity; |
19246 | 221 |
val axclass_instance_arity_i = |
19412 | 222 |
gen_instance (Logic.mk_arities oo Sign.cert_arity) AxClass.add_arity; |
20182
79c9ff40d760
tactic for prove_instance_arity now gets definition theorems as arguments
haftmann
parents:
20175
diff
changeset
|
223 |
val axclass_instance_sort = |
79c9ff40d760
tactic for prove_instance_arity now gets definition theorems as arguments
haftmann
parents:
20175
diff
changeset
|
224 |
gen_instance (single oo (Logic.mk_classrel oo AxClass.read_classrel)) AxClass.add_classrel I o single; |
19246 | 225 |
|
226 |
end; |
|
227 |
||
228 |
||
20455 | 229 |
|
230 |
(** classes and instances **) |
|
19038 | 231 |
|
232 |
local |
|
233 |
||
19957 | 234 |
fun add_axclass_i (name, supsort) params axs thy = |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
235 |
let |
19395 | 236 |
val (c, thy') = thy |
19957 | 237 |
|> AxClass.define_class_i (name, supsort) params axs; |
19518 | 238 |
val {intro, axioms, ...} = AxClass.get_definition thy' c; |
19395 | 239 |
in ((c, (intro, axioms)), thy') end; |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
240 |
|
20385 | 241 |
(*FIXME proper locale interface*) |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
242 |
fun prove_interpretation_i (prfx, atts) expr insts tac thy = |
19038 | 243 |
let |
20368
2461a0485623
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
wenzelm
parents:
20353
diff
changeset
|
244 |
fun ad_hoc_term (Const (c, ty)) = |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
245 |
let |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
246 |
val p = setmp show_types true (setmp show_sorts true (setmp print_mode [] (Sign.pretty_typ thy))) ty; |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
247 |
val s = c ^ "::" ^ Pretty.output p; |
20368
2461a0485623
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
wenzelm
parents:
20353
diff
changeset
|
248 |
in s end |
2461a0485623
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
wenzelm
parents:
20353
diff
changeset
|
249 |
| ad_hoc_term t = |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
250 |
let |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
251 |
val p = setmp show_types true (setmp show_sorts true (setmp print_mode [] (Sign.pretty_term thy))) t; |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
252 |
val s = Pretty.output p; |
20368
2461a0485623
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
wenzelm
parents:
20353
diff
changeset
|
253 |
in s end; |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
254 |
in |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
255 |
thy |
20368
2461a0485623
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
wenzelm
parents:
20353
diff
changeset
|
256 |
|> Locale.interpretation I (prfx, atts) expr (map (Option.map ad_hoc_term) insts) |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
257 |
|> Proof.global_terminal_proof (Method.Basic (fn _ => Method.SIMPLE_METHOD tac), NONE) |
20368
2461a0485623
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
wenzelm
parents:
20353
diff
changeset
|
258 |
|> ProofContext.theory_of |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
259 |
end; |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
260 |
|
19150 | 261 |
fun gen_class add_locale prep_class bname raw_supclasses raw_elems thy = |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
262 |
let |
20428 | 263 |
val (elems, exprs) = fold_rev (fn Locale.Elem e => apfst (cons e) |
264 |
| Locale.Expr e => apsnd (cons e)) |
|
265 |
raw_elems ([], []); |
|
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
266 |
val supclasses = map (prep_class thy) raw_supclasses; |
19038 | 267 |
val supsort = |
268 |
supclasses |
|
19966 | 269 |
|> map (#name_axclass o fst o the_class_data thy) |
19468 | 270 |
|> Sign.certify_sort thy |
19038 | 271 |
|> null ? K (Sign.defaultS thy); |
20428 | 272 |
val expr_supclasses = Locale.Merge |
273 |
(map (Locale.Locale o #name_locale o fst o the_class_data thy) supclasses); |
|
274 |
val expr = Locale.Merge (map (Locale.Locale o #name_locale o fst o the_class_data thy) supclasses |
|
275 |
@ exprs); |
|
19280
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
276 |
val mapp_sup = AList.make |
19468 | 277 |
(the o AList.lookup (op =) ((flat o map (the_parm_map thy) o the_ancestry thy) supclasses)) |
20428 | 278 |
((map (fst o fst) o Locale.parameters_of_expr thy) expr_supclasses); |
19038 | 279 |
fun extract_tyvar_consts thy name_locale = |
280 |
let |
|
20175 | 281 |
fun extract_classvar ((c, ty), _) w = |
20428 | 282 |
(case Term.add_tfreesT ty [] |
20175 | 283 |
of [(v, _)] => (case w |
20385 | 284 |
of SOME u => if u = v then w else error ("Additonal type variable in type signature of constant " ^ quote c) |
20175 | 285 |
| NONE => SOME v) |
20385 | 286 |
| [] => error ("No type variable in type signature of constant " ^ quote c) |
287 |
| _ => error ("More than one type variable in type signature of constant " ^ quote c)); |
|
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
288 |
val consts1 = |
19038 | 289 |
Locale.parameters_of thy name_locale |
290 |
|> map (apsnd Syntax.unlocalize_mixfix) |
|
20175 | 291 |
val SOME v = fold extract_classvar consts1 NONE; |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
292 |
val consts2 = map ((apfst o apsnd) (subst_clsvar v (TFree (v, [])))) consts1; |
19280
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
293 |
in (v, chop (length mapp_sup) consts2) end; |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
294 |
fun add_consts v raw_cs_sup raw_cs_this thy = |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
295 |
let |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
296 |
fun add_global_const ((c, ty), syn) thy = |
20265 | 297 |
((c, (Sign.full_name thy c, ty)), |
298 |
thy |
|
20703 | 299 |
|> Sign.add_consts_authentic [(c, ty |> subst_clsvar v (TFree (v, Sign.defaultS thy)), syn)]); |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
300 |
in |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
301 |
thy |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
302 |
|> fold_map add_global_const raw_cs_this |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
303 |
end; |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
304 |
fun extract_assumes thy name_locale cs_mapp = |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
305 |
let |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
306 |
val subst_assume = |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
307 |
map_aterms (fn Free (c, ty) => Const ((fst o the o AList.lookup (op =) cs_mapp) c, ty) |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
308 |
| t => t) |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
309 |
fun prep_asm ((name, atts), ts) = |
19957 | 310 |
((NameSpace.base name, map (Attrib.attribute thy) atts), map subst_assume ts) |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
311 |
in |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
312 |
(map prep_asm o Locale.local_asms_of thy) name_locale |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
313 |
end; |
19280
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
314 |
fun add_global_constraint v class (_, (c, ty)) thy = |
19038 | 315 |
thy |
19136 | 316 |
|> Sign.add_const_constraint_i (c, SOME (subst_clsvar v (TFree (v, [class])) ty)); |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
317 |
fun mk_const thy class v (c, ty) = |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
318 |
Const (c, subst_clsvar v (TFree (v, [class])) ty); |
19038 | 319 |
in |
320 |
thy |
|
20428 | 321 |
|> add_locale bname expr elems |
19928 | 322 |
|-> (fn (name_locale, ctxt) => |
19038 | 323 |
`(fn thy => extract_tyvar_consts thy name_locale) |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
324 |
#-> (fn (v, (raw_cs_sup, raw_cs_this)) => |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
325 |
add_consts v raw_cs_sup raw_cs_this |
19280
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
326 |
#-> (fn mapp_this => |
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
327 |
`(fn thy => extract_assumes thy name_locale (mapp_sup @ mapp_this)) |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
328 |
#-> (fn loc_axioms => |
19957 | 329 |
add_axclass_i (bname, supsort) (map (fst o snd) mapp_this) loc_axioms |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
330 |
#-> (fn (name_axclass, (_, ax_axioms)) => |
19280
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
331 |
fold (add_global_constraint v name_axclass) mapp_this |
19966 | 332 |
#> add_class_data (name_locale, (name_locale, name_axclass, v, mapp_this, |
333 |
map (fst o fst) loc_axioms)) |
|
20601 | 334 |
#> prove_interpretation_i (bname, []) |
19280
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
335 |
(Locale.Locale name_locale) (map (SOME o mk_const thy name_axclass v) (map snd (mapp_sup @ mapp_this))) |
19929 | 336 |
((ALLGOALS o ProofContext.fact_tac) ax_axioms) |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
337 |
#> pair ctxt |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
338 |
))))) |
19038 | 339 |
end; |
340 |
||
341 |
in |
|
342 |
||
20385 | 343 |
val class = gen_class (Locale.add_locale false) read_class; |
19928 | 344 |
val class_i = gen_class (Locale.add_locale_i false) certify_class; |
19038 | 345 |
|
20455 | 346 |
end; (*local*) |
19038 | 347 |
|
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
348 |
local |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
349 |
|
20175 | 350 |
fun gen_read_def thy prep_att read_def ((raw_name, raw_atts), raw_t) = |
19038 | 351 |
let |
19957 | 352 |
val (_, t) = read_def thy (raw_name, raw_t); |
353 |
val ((c, ty), _) = Sign.cert_def (Sign.pp thy) t; |
|
354 |
val atts = map (prep_att thy) raw_atts; |
|
20175 | 355 |
val insts = (Consts.typargs (Sign.consts_of thy) (c, ty)) |
19957 | 356 |
val name = case raw_name |
20385 | 357 |
of "" => NONE |
358 |
| _ => SOME raw_name; |
|
20175 | 359 |
in (c, (insts, ((name, t), atts))) end; |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
360 |
|
19957 | 361 |
fun read_def thy = gen_read_def thy Attrib.attribute read_axm; |
362 |
fun read_def_i thy = gen_read_def thy (K I) (K I); |
|
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
363 |
|
20175 | 364 |
fun gen_instance_arity prep_arity prep_att read_def do_proof raw_arities (raw_name, raw_atts) raw_defs theory = |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
365 |
let |
20175 | 366 |
fun prep_arity' ((tyco, asorts), sort) = prep_arity theory (tyco, asorts, sort); |
367 |
val arities = map prep_arity' raw_arities; |
|
368 |
val arities_pair = map (fn (tyco, asorts, sort) => ((tyco, asorts), sort)) arities; |
|
369 |
val _ = if null arities then error "at least one arity must be given" else (); |
|
370 |
val _ = case (duplicates (op =) o map #1) arities |
|
371 |
of [] => () |
|
372 |
| dupl_tycos => error ("type constructors occur more than once in arities: " |
|
373 |
^ (commas o map quote) dupl_tycos); |
|
20393 | 374 |
val (bind_always, name) = case raw_name |
375 |
of "" => (false, Thm.def_name ((space_implode "_" o map NameSpace.base) |
|
20175 | 376 |
(maps (fn (tyco, _, sort) => sort @ [tyco]) |
20393 | 377 |
(sort (fn ((tyco1, _, _), (tyco2, _, _)) => string_ord (tyco1, tyco2)) arities)))) |
378 |
| _ => (true, raw_name); |
|
19136 | 379 |
val atts = map (prep_att theory) raw_atts; |
20175 | 380 |
fun already_defined (c, ty_inst) = |
381 |
is_some (find_first (fn (_, { lhs = [ty_inst'], ...}) => |
|
382 |
Sign.typ_instance theory (ty_inst', ty_inst) orelse Sign.typ_instance theory (ty_inst, ty_inst')) |
|
383 |
(Defs.specifications_of (Theory.defs_of theory) c)); |
|
384 |
fun get_consts_class tyco ty class = |
|
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
385 |
let |
19966 | 386 |
val data = (fst o the_class_data theory) class; |
19957 | 387 |
val subst_ty = map_type_tfree (fn (v, sort) => |
20175 | 388 |
if #var data = v then ty else TVar ((v, 0), sort)); |
19957 | 389 |
in |
390 |
(map_filter (fn (_, (c, ty)) => |
|
20175 | 391 |
if already_defined (c, ty) |
392 |
then NONE else SOME ((c, ((tyco, class), subst_ty ty)))) o #consts) data |
|
19957 | 393 |
end; |
20175 | 394 |
fun get_consts_sort (tyco, asorts, sort) = |
19957 | 395 |
let |
20192
956cd30ef3be
renamed Name.give_names to Name.names and moved Name.alphanum to Symbol.alphanum
haftmann
parents:
20191
diff
changeset
|
396 |
val ty = Type (tyco, map (fn (v, sort) => TVar ((v, 0), sort)) (Name.names Name.context "'a" asorts)) |
20175 | 397 |
in maps (get_consts_class tyco ty) (the_ancestry theory sort) end; |
398 |
val cs = maps get_consts_sort arities; |
|
20385 | 399 |
fun mk_typnorm thy (ty, ty_sc) = |
400 |
case try (Sign.typ_match thy (Logic.varifyT ty_sc, ty)) Vartab.empty |
|
401 |
of SOME env => SOME (Logic.varifyT #> Envir.typ_subst_TVars env #> Logic.unvarifyT) |
|
402 |
| NONE => NONE; |
|
20175 | 403 |
fun read_defs defs cs thy_read = |
404 |
let |
|
19957 | 405 |
fun read raw_def cs = |
406 |
let |
|
20385 | 407 |
val (c, (inst, ((name_opt, t), atts))) = read_def thy_read raw_def; |
408 |
val ty = Consts.instance (Sign.consts_of thy_read) (c, inst); |
|
20175 | 409 |
val ((tyco, class), ty') = case AList.lookup (op =) cs c |
19957 | 410 |
of NONE => error ("superfluous definition for constant " ^ quote c) |
19966 | 411 |
| SOME class_ty => class_ty; |
20385 | 412 |
val name = case name_opt |
20628 | 413 |
of NONE => Thm.def_name (Logic.name_arity (tyco, [], c)) |
20385 | 414 |
| SOME name => name; |
415 |
val t' = case mk_typnorm thy_read (ty', ty) |
|
19966 | 416 |
of NONE => error ("superfluous definition for constant " ^ |
417 |
quote c ^ "::" ^ Sign.string_of_typ thy_read ty) |
|
20548
8ef25fe585a8
renamed Term.map_term_types to Term.map_types (cf. Term.fold_types);
wenzelm
parents:
20465
diff
changeset
|
418 |
| SOME norm => map_types norm t |
20175 | 419 |
in (((class, tyco), ((name, t'), atts)), AList.delete (op =) c cs) end; |
19957 | 420 |
in fold_map read defs cs end; |
20175 | 421 |
val (defs, _) = assume_arities_thy theory arities_pair (read_defs raw_defs cs); |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
422 |
fun get_remove_contraint c thy = |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
423 |
let |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
424 |
val ty = Sign.the_const_constraint thy c; |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
425 |
in |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
426 |
thy |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
427 |
|> Sign.add_const_constraint_i (c, NONE) |
20385 | 428 |
|> pair (c, Logic.unvarifyT ty) |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
429 |
end; |
19966 | 430 |
fun add_defs defs thy = |
431 |
thy |
|
432 |
|> PureThy.add_defs_i true (map snd defs) |
|
433 |
|-> (fn thms => pair (map fst defs ~~ thms)); |
|
434 |
fun note_all thy = |
|
20175 | 435 |
(*FIXME: should avoid binding duplicated names*) |
19966 | 436 |
let |
20393 | 437 |
val bind = bind_always orelse not (can (PureThy.get_thms thy) (Name name)); |
20175 | 438 |
val thms = maps (fn (tyco, _, sort) => maps (fn class => |
439 |
Symtab.lookup_list |
|
440 |
((snd o the_class_data thy) class) tyco) (the_ancestry thy sort)) arities; |
|
20393 | 441 |
in if bind then |
19966 | 442 |
thy |
20601 | 443 |
|> PureThy.note_thmss_i (*qualified*) PureThy.internalK [((name, atts), [(thms, [])])] |
19966 | 444 |
|> snd |
20393 | 445 |
else |
446 |
thy |
|
19966 | 447 |
end; |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
448 |
fun after_qed cs thy = |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
449 |
thy |
19412 | 450 |
|> fold Sign.add_const_constraint_i (map (apsnd SOME) cs); |
19038 | 451 |
in |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
452 |
theory |
20385 | 453 |
|> fold_map get_remove_contraint (map fst cs |> distinct (op =)) |
19966 | 454 |
||>> add_defs defs |
20076 | 455 |
|-> (fn (cs, def_thms) => |
20175 | 456 |
fold add_inst_def def_thms |
19966 | 457 |
#> note_all |
20182
79c9ff40d760
tactic for prove_instance_arity now gets definition theorems as arguments
haftmann
parents:
20175
diff
changeset
|
458 |
#> do_proof (map snd def_thms) (after_qed cs) arities) |
19038 | 459 |
end; |
460 |
||
19957 | 461 |
fun instance_arity' do_proof = gen_instance_arity Sign.read_arity Attrib.attribute |
462 |
read_def do_proof; |
|
463 |
fun instance_arity_i' do_proof = gen_instance_arity Sign.cert_arity (K I) |
|
464 |
read_def_i do_proof; |
|
20182
79c9ff40d760
tactic for prove_instance_arity now gets definition theorems as arguments
haftmann
parents:
20175
diff
changeset
|
465 |
fun tactic_proof tac def_thms after_qed arities = |
79c9ff40d760
tactic for prove_instance_arity now gets definition theorems as arguments
haftmann
parents:
20175
diff
changeset
|
466 |
fold (fn arity => AxClass.prove_arity arity (tac def_thms)) arities |
20175 | 467 |
#> after_qed; |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
468 |
|
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
469 |
in |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
470 |
|
20182
79c9ff40d760
tactic for prove_instance_arity now gets definition theorems as arguments
haftmann
parents:
20175
diff
changeset
|
471 |
val instance_arity = instance_arity' (K axclass_instance_arity_i); |
79c9ff40d760
tactic for prove_instance_arity now gets definition theorems as arguments
haftmann
parents:
20175
diff
changeset
|
472 |
val instance_arity_i = instance_arity_i' (K axclass_instance_arity_i); |
19150 | 473 |
val prove_instance_arity = instance_arity_i' o tactic_proof; |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
474 |
|
20455 | 475 |
end; (*local*) |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
476 |
|
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
477 |
local |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
478 |
|
20368
2461a0485623
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
wenzelm
parents:
20353
diff
changeset
|
479 |
fun prove_interpretation_in tac after_qed (name, expr) thy = |
19136 | 480 |
thy |
20368
2461a0485623
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
wenzelm
parents:
20353
diff
changeset
|
481 |
|> Locale.interpretation_in_locale (ProofContext.theory after_qed) (name, expr) |
19136 | 482 |
|> Proof.global_terminal_proof (Method.Basic (fn _ => Method.SIMPLE_METHOD tac), NONE) |
20368
2461a0485623
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
wenzelm
parents:
20353
diff
changeset
|
483 |
|> ProofContext.theory_of; |
19038 | 484 |
|
20385 | 485 |
(*FIXME very ad-hoc, needs proper locale interface*) |
19150 | 486 |
fun gen_instance_sort prep_class prep_sort do_proof (raw_class, raw_sort) theory = |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
487 |
let |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
488 |
val class = prep_class theory raw_class; |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
489 |
val sort = prep_sort theory raw_sort; |
19966 | 490 |
val loc_name = (#name_locale o fst o the_class_data theory) class; |
19468 | 491 |
val loc_expr = |
19966 | 492 |
(Locale.Merge o map (Locale.Locale o #name_locale o fst o the_class_data theory)) sort; |
493 |
val const_names = (map (NameSpace.base o fst o snd) |
|
494 |
o maps (#consts o fst o the_class_data theory) |
|
495 |
o the_ancestry theory) [class]; |
|
20106 | 496 |
fun get_thms thy = |
497 |
the_ancestry thy sort |
|
498 |
|> AList.make (the_propnames thy) |
|
499 |
|> map (apsnd (map (NameSpace.append (loc_name)))) |
|
500 |
|> map_filter (fn (superclass, thm_names) => |
|
501 |
case map_filter (try (PureThy.get_thm thy o Name)) thm_names |
|
502 |
of [] => NONE |
|
503 |
| thms => SOME (superclass, thms)); |
|
504 |
fun after_qed thy = |
|
505 |
thy |
|
506 |
|> `get_thms |
|
507 |
|-> fold (fn (supclass, thms) => I |
|
508 |
AxClass.prove_classrel (class, supclass) |
|
509 |
(ALLGOALS (K (intro_classes_tac [])) THEN |
|
510 |
(ALLGOALS o ProofContext.fact_tac) thms)) |
|
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
511 |
in |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
512 |
theory |
19136 | 513 |
|> do_proof after_qed (loc_name, loc_expr) |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
514 |
end; |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
515 |
|
20385 | 516 |
fun instance_sort' do_proof = gen_instance_sort read_class read_sort do_proof; |
19150 | 517 |
fun instance_sort_i' do_proof = gen_instance_sort certify_class certify_sort do_proof; |
20368
2461a0485623
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
wenzelm
parents:
20353
diff
changeset
|
518 |
val setup_proof = Locale.interpretation_in_locale o ProofContext.theory; |
19136 | 519 |
val tactic_proof = prove_interpretation_in; |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
520 |
|
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
521 |
in |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
522 |
|
19150 | 523 |
val instance_sort = instance_sort' setup_proof; |
524 |
val instance_sort_i = instance_sort_i' setup_proof; |
|
525 |
val prove_instance_sort = instance_sort_i' o tactic_proof; |
|
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
526 |
|
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
527 |
end; (* local *) |
19038 | 528 |
|
19957 | 529 |
|
18168 | 530 |
|
20455 | 531 |
(** code generation view **) |
18168 | 532 |
|
20455 | 533 |
fun operational_algebra thy = |
534 |
Sorts.project_algebra (Sign.pp thy) |
|
535 |
(is_operational_class thy) (Sign.classes_of thy); |
|
18168 | 536 |
|
20455 | 537 |
fun the_consts_sign thy class = |
538 |
let |
|
539 |
val _ = if is_operational_class thy class then () else error ("no operational class: " ^ quote class); |
|
540 |
val data = (fst o the_class_data thy) class |
|
541 |
in (#var data, (map snd o #consts) data) end; |
|
19213 | 542 |
|
20455 | 543 |
fun the_inst_sign thy (class, tyco) = |
18168 | 544 |
let |
20455 | 545 |
val _ = if is_operational_class thy class then () else error ("no operational class: " ^ quote class); |
546 |
val asorts = Sign.arity_sorts thy tyco [class]; |
|
547 |
val (clsvar, const_sign) = the_consts_sign thy class; |
|
548 |
fun add_var sort used = |
|
549 |
let val v = hd (Name.invents used "'a" 1); |
|
550 |
in ((v, sort), Name.declare v used) end; |
|
551 |
val (vsorts, _) = |
|
552 |
Name.context |
|
553 |
|> Name.declare clsvar |
|
554 |
|> fold (fn (_, ty) => fold Name.declare |
|
555 |
((map (fst o fst) o typ_tvars) ty @ map fst (Term.add_tfreesT ty []))) const_sign |
|
556 |
|> fold_map add_var asorts; |
|
557 |
val ty_inst = Type (tyco, map TFree vsorts); |
|
558 |
val inst_signs = map (apsnd (subst_clsvar clsvar ty_inst)) const_sign; |
|
559 |
in (vsorts, inst_signs) end; |
|
18168 | 560 |
|
18884 | 561 |
|
20455 | 562 |
|
563 |
(** toplevel interface **) |
|
18515 | 564 |
|
565 |
local |
|
566 |
||
567 |
structure P = OuterParse |
|
568 |
and K = OuterKeyword |
|
569 |
||
570 |
in |
|
571 |
||
20385 | 572 |
val (classK, instanceK, print_classesK) = ("class", "instance", "print_classes") |
18515 | 573 |
|
20385 | 574 |
fun wrap_add_instance_sort (class, sort) thy = |
575 |
(if forall (is_some o lookup_class_data thy) (Sign.read_sort thy sort) |
|
576 |
then instance_sort else axclass_instance_sort) (class, sort) thy; |
|
19136 | 577 |
|
20385 | 578 |
val class_subP = P.name -- Scan.repeat (P.$$$ "+" |-- P.name) >> (op ::); |
20428 | 579 |
val class_bodyP = P.!!! (Scan.repeat1 P.locale_element); |
20385 | 580 |
|
20601 | 581 |
val parse_arity = |
582 |
(P.xname --| P.$$$ "::" -- P.!!! P.arity) |
|
19136 | 583 |
>> (fn (tyco, (asorts, sort)) => ((tyco, asorts), sort)); |
18911 | 584 |
|
18515 | 585 |
val classP = |
20385 | 586 |
OuterSyntax.command classK "define operational type class" K.thy_decl ( |
18849
05a16861d3a5
added three times overloaded Isar instance command
haftmann
parents:
18829
diff
changeset
|
587 |
P.name --| P.$$$ "=" |
19280
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
588 |
-- ( |
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
589 |
class_subP --| P.$$$ "+" -- class_bodyP |
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
590 |
|| class_subP >> rpair [] |
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
591 |
|| class_bodyP >> pair [] |
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
592 |
) >> (Toplevel.theory_context |
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
593 |
o (fn (bname, (supclasses, elems)) => class bname supclasses elems))); |
18515 | 594 |
|
18575 | 595 |
val instanceP = |
18849
05a16861d3a5
added three times overloaded Isar instance command
haftmann
parents:
18829
diff
changeset
|
596 |
OuterSyntax.command instanceK "prove type arity or subclass relation" K.thy_goal (( |
20385 | 597 |
P.xname -- ((P.$$$ "\\<subseteq>" || P.$$$ "<") |-- P.!!! P.xname) >> wrap_add_instance_sort |
20601 | 598 |
|| P.opt_thm_name ":" -- (P.and_list1 parse_arity -- Scan.repeat (P.opt_thm_name ":" -- P.prop)) |
20175 | 599 |
>> (fn (("", []), ([((tyco, asorts), sort)], [])) => axclass_instance_arity I [(tyco, asorts, sort)] |
600 |
| (natts, (arities, defs)) => instance_arity arities natts defs) |
|
18849
05a16861d3a5
added three times overloaded Isar instance command
haftmann
parents:
18829
diff
changeset
|
601 |
) >> (Toplevel.print oo Toplevel.theory_to_proof)); |
18575 | 602 |
|
20385 | 603 |
val print_classesP = |
604 |
OuterSyntax.improper_command print_classesK "print classes of this theory" K.diag |
|
605 |
(Scan.succeed (Toplevel.no_timing o Toplevel.unknown_theory |
|
606 |
o Toplevel.keep (print_classes o Toplevel.theory_of))); |
|
607 |
||
608 |
val _ = OuterSyntax.add_parsers [classP, instanceP, print_classesP]; |
|
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
609 |
|
20455 | 610 |
end; (*local*) |
18515 | 611 |
|
20455 | 612 |
end; |