author | wenzelm |
Thu, 27 Apr 2006 15:06:35 +0200 | |
changeset 19482 | 9f11af8f7ef9 |
parent 19468 | 0afdd5023bfc |
child 19518 | 5204c73a4d46 |
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 |
|
19150 | 10 |
val class: bstring -> class list -> Element.context list -> theory |
18515 | 11 |
-> ProofContext.context * theory |
19150 | 12 |
val class_i: bstring -> class list -> Element.context_i list -> theory |
18515 | 13 |
-> ProofContext.context * theory |
19150 | 14 |
val instance_arity: (xstring * string list) * string |
19136 | 15 |
-> bstring * Attrib.src list -> ((bstring * Attrib.src list) * string) list |
18575 | 16 |
-> theory -> Proof.state |
19150 | 17 |
val instance_arity_i: (string * sort list) * sort |
19136 | 18 |
-> bstring * attribute list -> ((bstring * attribute list) * term) list |
18575 | 19 |
-> theory -> Proof.state |
19150 | 20 |
val prove_instance_arity: tactic -> (string * sort list) * sort |
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 |
|
19253 | 27 |
val use_cp_instance: bool ref; |
28 |
||
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
29 |
val intern_class: theory -> xstring -> class |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
30 |
val intern_sort: theory -> sort -> sort |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
31 |
val extern_class: theory -> class -> xstring |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
32 |
val extern_sort: theory -> sort -> sort |
18755 | 33 |
val certify_class: theory -> class -> class |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
34 |
val certify_sort: theory -> sort -> sort |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
35 |
val read_sort: theory -> string -> sort |
18884 | 36 |
val operational_sort_of: theory -> sort -> sort |
18702 | 37 |
val the_superclasses: theory -> class -> class list |
38 |
val the_consts_sign: theory -> class -> string * (string * typ) list |
|
18168 | 39 |
val lookup_const_class: theory -> string -> class option |
19213 | 40 |
val the_instances: theory -> class -> (string * ((sort list) * string)) list |
18702 | 41 |
val the_inst_sign: theory -> class * string -> (string * sort) list * (string * typ) list |
19213 | 42 |
val get_classtab: theory -> (string * string) list Symtab.table |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
43 |
|
18702 | 44 |
val print_classes: theory -> unit |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
45 |
val intro_classes_tac: thm list -> tactic |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
46 |
val default_intro_classes_tac: thm list -> tactic |
18168 | 47 |
|
48 |
type sortcontext = (string * sort) list |
|
18884 | 49 |
datatype classlookup = Instance of (class * string) * classlookup list list |
19253 | 50 |
| Lookup of class list * (string * (int * int)) |
18168 | 51 |
val extract_sortctxt: theory -> typ -> sortcontext |
18884 | 52 |
val extract_classlookup: theory -> string * typ -> classlookup list list |
53 |
val extract_classlookup_inst: theory -> class * string -> class -> classlookup list list |
|
19253 | 54 |
val extract_classlookup_member: theory -> typ * typ -> classlookup list list |
18168 | 55 |
end; |
56 |
||
57 |
structure ClassPackage: CLASS_PACKAGE = |
|
58 |
struct |
|
59 |
||
60 |
||
18708 | 61 |
(* theory data *) |
18168 | 62 |
|
19456 | 63 |
datatype class_data = ClassData of { |
18575 | 64 |
name_locale: string, |
65 |
name_axclass: string, |
|
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
66 |
intro: thm option, |
18575 | 67 |
var: string, |
19280
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
68 |
consts: (string * (string * typ)) list |
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
69 |
(*locale parameter ~> toplevel const*) |
18168 | 70 |
}; |
71 |
||
19456 | 72 |
fun rep_classdata (ClassData c) = c; |
73 |
fun eq_classdata (ClassData { |
|
74 |
name_locale = name_locale1, name_axclass = name_axclass1, intro = intro1, |
|
75 |
var = var1, consts = consts1}, ClassData { |
|
76 |
name_locale = name_locale2, name_axclass = name_axclass2, intro = intro2, |
|
77 |
var = var2, consts = consts2}) = |
|
78 |
name_locale1 = name_locale2 andalso name_axclass1 = name_axclass2 |
|
79 |
andalso eq_opt eq_thm (intro1, intro2) andalso var1 = var2 |
|
80 |
andalso eq_list (eq_pair (op =) (eq_pair (op =) (Type.eq_type Vartab.empty))) |
|
81 |
(consts1, consts2); |
|
82 |
||
18575 | 83 |
structure ClassData = TheoryDataFun ( |
18168 | 84 |
struct |
85 |
val name = "Pure/classes"; |
|
19213 | 86 |
type T = (class_data Graph.T |
87 |
* (string * (sort list * string)) list Symtab.table) |
|
19456 | 88 |
(*class ~> tyco ~> (arity, thyname)*) |
19213 | 89 |
* class Symtab.table; |
90 |
val empty = ((Graph.empty, Symtab.empty), Symtab.empty); |
|
18168 | 91 |
val copy = I; |
92 |
val extend = I; |
|
19456 | 93 |
fun merge _ (((g1, c1), f1) : T, ((g2, c2), f2)) = |
94 |
((Graph.merge eq_classdata (g1, g2), Symtab.join (fn _ => AList.merge (op =) (op =)) (c1, c2)), |
|
19213 | 95 |
Symtab.merge (op =) (f1, f2)); |
96 |
fun print thy ((gr, _), _) = |
|
18575 | 97 |
let |
19456 | 98 |
fun pretty_class gr (name, ClassData {name_locale, name_axclass, intro, var, consts}) = |
18575 | 99 |
(Pretty.block o Pretty.fbreaks) [ |
100 |
Pretty.str ("class " ^ name ^ ":"), |
|
101 |
(Pretty.block o Pretty.fbreaks) ( |
|
102 |
Pretty.str "superclasses: " |
|
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
103 |
:: (map Pretty.str o Graph.imm_succs gr) name |
18575 | 104 |
), |
105 |
Pretty.str ("locale: " ^ name_locale), |
|
106 |
Pretty.str ("axclass: " ^ name_axclass), |
|
107 |
Pretty.str ("class variable: " ^ var), |
|
108 |
(Pretty.block o Pretty.fbreaks) ( |
|
109 |
Pretty.str "constants: " |
|
19280
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
110 |
:: map (fn (_, (c, ty)) => Pretty.str (c ^ " :: " ^ Sign.string_of_typ thy ty)) consts |
18575 | 111 |
) |
112 |
] |
|
113 |
in |
|
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
114 |
(Pretty.writeln o Pretty.chunks o map (pretty_class gr) |
19468 | 115 |
o AList.make (Graph.get_node gr) o flat o Graph.strong_conn) gr |
18575 | 116 |
end; |
18168 | 117 |
end |
118 |
); |
|
119 |
||
18708 | 120 |
val _ = Context.add_setup ClassData.init; |
18575 | 121 |
val print_classes = ClassData.print; |
122 |
||
19038 | 123 |
|
124 |
(* queries *) |
|
125 |
||
19456 | 126 |
val lookup_class_data = Option.map rep_classdata oo try o Graph.get_node o fst o fst o ClassData.get; |
19213 | 127 |
val the_instances = these oo Symtab.lookup o snd o fst o ClassData.get; |
128 |
val lookup_const_class = Symtab.lookup o snd o ClassData.get; |
|
18168 | 129 |
|
19038 | 130 |
fun the_class_data thy class = |
18168 | 131 |
case lookup_class_data thy class |
18755 | 132 |
of NONE => error ("undeclared operational class " ^ quote class) |
18168 | 133 |
| SOME data => data; |
134 |
||
19280
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
135 |
val is_class = is_some oo lookup_class_data; |
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
136 |
|
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
137 |
fun is_operational_class thy cls = |
18702 | 138 |
lookup_class_data thy cls |
139 |
|> Option.map (not o null o #consts) |
|
140 |
|> the_default false; |
|
18168 | 141 |
|
18884 | 142 |
fun operational_sort_of thy sort = |
18360 | 143 |
let |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
144 |
fun get_sort class = |
19280
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
145 |
if is_operational_class thy class |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
146 |
then [class] |
19412 | 147 |
else operational_sort_of thy (Sign.super_classes thy class); |
18360 | 148 |
in |
149 |
map get_sort sort |
|
19468 | 150 |
|> flat |
19412 | 151 |
|> Sign.certify_sort thy |
18360 | 152 |
end; |
18168 | 153 |
|
18702 | 154 |
fun the_superclasses thy class = |
18515 | 155 |
if is_class thy class |
156 |
then |
|
19412 | 157 |
Sign.super_classes thy class |
18884 | 158 |
|> operational_sort_of thy |
18515 | 159 |
else |
19280
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
160 |
error ("no class: " ^ class); |
18168 | 161 |
|
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
162 |
fun get_superclass_derivation thy (subclass, superclass) = |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
163 |
if subclass = superclass |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
164 |
then SOME [subclass] |
19213 | 165 |
else case Graph.find_paths ((fst o fst o ClassData.get) thy) (subclass, superclass) |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
166 |
of [] => NONE |
19280
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
167 |
| (p::_) => (SOME o filter (is_operational_class thy)) p; |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
168 |
|
19038 | 169 |
fun the_ancestry thy classes = |
170 |
let |
|
171 |
fun ancestry class anc = |
|
172 |
anc |
|
173 |
|> cons class |
|
174 |
|> fold ancestry (the_superclasses thy class); |
|
175 |
in fold ancestry classes [] end; |
|
176 |
||
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
177 |
fun the_intros thy = |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
178 |
let |
19213 | 179 |
val gr = (fst o fst o ClassData.get) thy; |
19482
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
wenzelm
parents:
19468
diff
changeset
|
180 |
in (map_filter (#intro o rep_classdata o Graph.get_node gr) o Graph.keys) gr end; |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
181 |
|
19038 | 182 |
fun subst_clsvar v ty_subst = |
183 |
map_type_tfree (fn u as (w, _) => |
|
184 |
if w = v then ty_subst else TFree u); |
|
185 |
||
19280
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
186 |
fun the_parm_map thy class = |
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
187 |
let |
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
188 |
val data = the_class_data thy class |
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
189 |
in (#consts data) end; |
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
190 |
|
18702 | 191 |
fun the_consts_sign thy class = |
18168 | 192 |
let |
19038 | 193 |
val data = the_class_data thy class |
19280
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
194 |
in (#var data, (map snd o #consts) data) end; |
18702 | 195 |
|
196 |
fun the_inst_sign thy (class, tyco) = |
|
18168 | 197 |
let |
19280
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
198 |
val _ = if is_operational_class thy class then () else error ("no operational class: " ^ class); |
19468 | 199 |
val arity = Sign.arity_sorts thy tyco [class]; |
19038 | 200 |
val clsvar = (#var o the_class_data thy) class; |
18702 | 201 |
val const_sign = (snd o the_consts_sign thy) class; |
202 |
fun add_var sort used = |
|
203 |
let |
|
204 |
val v = hd (Term.invent_names used "'a" 1) |
|
205 |
in ((v, sort), v::used) end; |
|
206 |
val (vsorts, _) = |
|
19213 | 207 |
[clsvar] |
18702 | 208 |
|> fold (fn (_, ty) => curry (gen_union (op =)) |
209 |
((map (fst o fst) o typ_tvars) ty @ (map fst o typ_tfrees) ty)) const_sign |
|
210 |
|> fold_map add_var arity; |
|
211 |
val ty_inst = Type (tyco, map (fn (v, sort) => TVar ((v, 0), sort)) vsorts); |
|
212 |
val inst_signs = map (apsnd (subst_clsvar clsvar ty_inst)) const_sign; |
|
213 |
in (vsorts, inst_signs) end; |
|
18168 | 214 |
|
215 |
fun get_classtab thy = |
|
19213 | 216 |
(Symtab.map o map) |
217 |
(fn (tyco, (_, thyname)) => (tyco, thyname)) ((snd o fst o ClassData.get) thy); |
|
18168 | 218 |
|
219 |
||
19038 | 220 |
(* updaters *) |
221 |
||
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
222 |
fun add_class_data (class, (superclasses, name_locale, name_axclass, intro, var, consts)) = |
19213 | 223 |
ClassData.map (fn ((gr, tab), consttab) => (( |
224 |
gr |
|
19456 | 225 |
|> Graph.new_node (class, ClassData { |
19038 | 226 |
name_locale = name_locale, |
227 |
name_axclass = name_axclass, |
|
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
228 |
intro = intro, |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
229 |
var = var, |
19213 | 230 |
consts = consts |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
231 |
}) |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
232 |
|> fold (curry Graph.add_edge_acyclic class) superclasses, |
19213 | 233 |
tab |
234 |
|> Symtab.update (class, [])), |
|
235 |
consttab |
|
19280
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
236 |
|> fold (fn (_, (c, _)) => Symtab.update (c, class)) consts |
19038 | 237 |
)); |
238 |
||
239 |
fun add_inst_data (class, inst) = |
|
19213 | 240 |
ClassData.map (fn ((gr, tab), consttab) => |
19280
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
241 |
let |
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
242 |
val undef_supclasses = class :: (filter (Symtab.defined tab) (Graph.all_succs gr [class])); |
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
243 |
in |
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
244 |
((gr, tab |> fold (fn class => Symtab.map_entry class (AList.update (op =) inst)) undef_supclasses), consttab) |
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
245 |
end); |
19038 | 246 |
|
247 |
||
248 |
(* name handling *) |
|
249 |
||
250 |
fun certify_class thy class = |
|
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
251 |
(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
|
252 |
|
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
253 |
fun certify_sort thy sort = |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
254 |
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
|
255 |
|
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
256 |
fun intern_class thy = |
19280
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
257 |
certify_class thy o Sign.intern_class thy; |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
258 |
|
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
259 |
fun intern_sort thy = |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
260 |
certify_sort thy o Sign.intern_sort thy; |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
261 |
|
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
262 |
fun extern_class thy = |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
263 |
Sign.extern_class thy o certify_class thy; |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
264 |
|
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
265 |
fun extern_sort thy = |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
266 |
Sign.extern_sort thy o certify_sort thy; |
19038 | 267 |
|
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
268 |
fun read_sort thy = |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
269 |
certify_sort thy o Sign.read_sort thy; |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
270 |
|
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
271 |
|
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
272 |
(* tactics and methods *) |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
273 |
|
19123 | 274 |
fun class_intros thy = |
275 |
AxClass.class_intros thy @ the_intros thy; |
|
19038 | 276 |
|
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
277 |
fun intro_classes_tac facts st = |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
278 |
(ALLGOALS (Method.insert_tac facts THEN' |
19123 | 279 |
REPEAT_ALL_NEW (resolve_tac (class_intros (Thm.theory_of_thm st)))) |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
280 |
THEN Tactic.distinct_subgoals_tac) st; |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
281 |
|
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
282 |
fun default_intro_classes_tac [] = intro_classes_tac [] |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
283 |
| default_intro_classes_tac _ = Tactical.no_tac; (*no error message!*) |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
284 |
|
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
285 |
fun default_tac rules ctxt facts = |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
286 |
HEADGOAL (Method.some_rule_tac rules ctxt facts) ORELSE |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
287 |
default_intro_classes_tac facts; |
19038 | 288 |
|
19468 | 289 |
val _ = Context.add_setup (Method.add_methods |
290 |
[("intro_classes", Method.no_args (Method.METHOD intro_classes_tac), |
|
291 |
"back-chain introduction rules of classes"), |
|
292 |
("default", Method.thms_ctxt_args (Method.METHOD oo default_tac), |
|
293 |
"apply some intro/elim rule")]); |
|
294 |
||
19038 | 295 |
|
19246 | 296 |
(* axclass instances *) |
297 |
||
298 |
local |
|
299 |
||
19412 | 300 |
fun gen_instance mk_prop add_thm after_qed inst thy = |
19246 | 301 |
thy |
302 |
|> ProofContext.init |
|
19412 | 303 |
|> Proof.theorem_i PureThy.internalK NONE (after_qed oo (fold o fold) add_thm) NONE ("", []) |
19246 | 304 |
(map (fn t => (("", []), [(t, ([], []))])) (mk_prop thy inst)); |
305 |
||
306 |
in |
|
307 |
||
308 |
val axclass_instance_subclass = |
|
19412 | 309 |
gen_instance (single oo (Logic.mk_classrel oo AxClass.read_classrel)) AxClass.add_classrel I; |
19246 | 310 |
val axclass_instance_arity = |
19412 | 311 |
gen_instance (Logic.mk_arities oo Sign.read_arity) AxClass.add_arity; |
19246 | 312 |
val axclass_instance_arity_i = |
19412 | 313 |
gen_instance (Logic.mk_arities oo Sign.cert_arity) AxClass.add_arity; |
19246 | 314 |
|
315 |
end; |
|
316 |
||
317 |
||
19038 | 318 |
(* classes and instances *) |
319 |
||
320 |
local |
|
321 |
||
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
322 |
fun intro_incr thy name expr = |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
323 |
let |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
324 |
fun fish_thm basename = |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
325 |
try (PureThy.get_thm thy) ((Name o NameSpace.append basename) "intro"); |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
326 |
in if expr = Locale.empty |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
327 |
then fish_thm name |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
328 |
else fish_thm (name ^ "_axioms") |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
329 |
end; |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
330 |
|
19468 | 331 |
fun add_locale name expr body thy = |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
332 |
thy |
19468 | 333 |
|> Locale.add_locale true name expr body |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
334 |
||>> `(fn thy => intro_incr thy name expr) |
19340 | 335 |
|-> (fn ((name, ctxt), intro) => pair ((name, intro), ctxt)); |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
336 |
|
19468 | 337 |
fun add_locale_i name expr body thy = |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
338 |
thy |
19468 | 339 |
|> Locale.add_locale_i true name expr body |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
340 |
||>> `(fn thy => intro_incr thy name expr) |
19340 | 341 |
|-> (fn ((name, ctxt), intro) => pair ((name, intro), ctxt)); |
19038 | 342 |
|
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
343 |
fun add_axclass_i (name, supsort) axs thy = |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
344 |
let |
19395 | 345 |
val (c, thy') = thy |
19434 | 346 |
|> AxClass.add_axclass_i (name, supsort) [] axs; |
19395 | 347 |
val {intro, axioms, ...} = AxClass.get_info thy' c; |
348 |
in ((c, (intro, axioms)), thy') end; |
|
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
349 |
|
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
350 |
fun prove_interpretation_i (prfx, atts) expr insts tac thy = |
19038 | 351 |
let |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
352 |
fun ad_hoc_term NONE = NONE |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
353 |
| ad_hoc_term (SOME (Const (c, ty))) = |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
354 |
let |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
355 |
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
|
356 |
val s = c ^ "::" ^ Pretty.output p; |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
357 |
val _ = writeln s; |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
358 |
in SOME s end |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
359 |
| ad_hoc_term (SOME t) = |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
360 |
let |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
361 |
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
|
362 |
val s = Pretty.output p; |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
363 |
val _ = writeln s; |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
364 |
in SOME s end; |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
365 |
in |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
366 |
thy |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
367 |
|> Locale.interpretation (prfx, atts) expr (map ad_hoc_term insts) |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
368 |
|> Proof.global_terminal_proof (Method.Basic (fn _ => Method.SIMPLE_METHOD tac), NONE) |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
369 |
|-> (fn _ => I) |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
370 |
end; |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
371 |
|
19150 | 372 |
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
|
373 |
let |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
374 |
val supclasses = map (prep_class thy) raw_supclasses; |
19038 | 375 |
val supsort = |
376 |
supclasses |
|
377 |
|> map (#name_axclass o the_class_data thy) |
|
19468 | 378 |
|> Sign.certify_sort thy |
19038 | 379 |
|> null ? K (Sign.defaultS thy); |
19468 | 380 |
val expr = (Locale.Merge o map (Locale.Locale o #name_locale o the_class_data thy)) supclasses; |
19280
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
381 |
val mapp_sup = AList.make |
19468 | 382 |
(the o AList.lookup (op =) ((flat o map (the_parm_map thy) o the_ancestry thy) supclasses)) |
19280
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
383 |
((map (fst o fst) o Locale.parameters_of_expr thy) expr); |
19038 | 384 |
fun extract_tyvar_consts thy name_locale = |
385 |
let |
|
386 |
fun extract_tyvar_name thy tys = |
|
387 |
fold (curry add_typ_tfrees) tys [] |
|
388 |
|> (fn [(v, sort)] => |
|
19468 | 389 |
if Sign.subsort thy (supsort, sort) |
19038 | 390 |
then v |
391 |
else error ("illegal sort constraint on class type variable: " ^ Sign.string_of_sort thy sort) |
|
392 |
| [] => error ("no class type variable") |
|
393 |
| vs => error ("more than one type variable: " ^ (commas o map (Sign.string_of_typ thy o TFree)) vs)) |
|
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
394 |
val consts1 = |
19038 | 395 |
Locale.parameters_of thy name_locale |
396 |
|> map (apsnd Syntax.unlocalize_mixfix) |
|
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
397 |
val v = (extract_tyvar_name thy o map (snd o fst)) consts1; |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
398 |
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
|
399 |
in (v, chop (length mapp_sup) consts2) end; |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
400 |
fun add_consts v raw_cs_sup raw_cs_this thy = |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
401 |
let |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
402 |
fun add_global_const ((c, ty), syn) thy = |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
403 |
thy |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
404 |
|> Sign.add_consts_i [(c, ty |> subst_clsvar v (TFree (v, Sign.defaultS thy)), syn)] |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
405 |
|> `(fn thy => (c, (Sign.intern_const thy c, ty))) |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
406 |
in |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
407 |
thy |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
408 |
|> fold_map add_global_const raw_cs_this |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
409 |
end; |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
410 |
fun extract_assumes thy name_locale cs_mapp = |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
411 |
let |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
412 |
val subst_assume = |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
413 |
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
|
414 |
| t => t) |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
415 |
fun prep_asm ((name, atts), ts) = |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
416 |
((name, map (Attrib.attribute thy) atts), map subst_assume ts) |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
417 |
in |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
418 |
(map prep_asm o Locale.local_asms_of thy) name_locale |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
419 |
end; |
19280
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
420 |
fun add_global_constraint v class (_, (c, ty)) thy = |
19038 | 421 |
thy |
19136 | 422 |
|> 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
|
423 |
fun mk_const thy class v (c, ty) = |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
424 |
Const (c, subst_clsvar v (TFree (v, [class])) ty); |
19038 | 425 |
in |
426 |
thy |
|
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
427 |
|> add_locale bname expr raw_elems |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
428 |
|-> (fn ((name_locale, intro), ctxt) => |
19038 | 429 |
`(fn thy => extract_tyvar_consts thy name_locale) |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
430 |
#-> (fn (v, (raw_cs_sup, raw_cs_this)) => |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
431 |
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
|
432 |
#-> (fn mapp_this => |
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
433 |
`(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
|
434 |
#-> (fn loc_axioms => |
19434 | 435 |
add_axclass_i (bname, supsort) (map (apfst (apfst (K ""))) loc_axioms) |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
436 |
#-> (fn (name_axclass, (_, ax_axioms)) => |
19280
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
437 |
fold (add_global_constraint v name_axclass) mapp_this |
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
438 |
#> add_class_data (name_locale, (supclasses, name_locale, name_axclass, intro, v, mapp_this)) |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
439 |
#> prove_interpretation_i (NameSpace.base name_locale, []) |
19280
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
440 |
(Locale.Locale name_locale) (map (SOME o mk_const thy name_axclass v) (map snd (mapp_sup @ mapp_this))) |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
441 |
((ALLGOALS o resolve_tac) ax_axioms) |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
442 |
#> pair ctxt |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
443 |
))))) |
19038 | 444 |
end; |
445 |
||
446 |
in |
|
447 |
||
19468 | 448 |
val class = gen_class add_locale intern_class; |
449 |
val class_i = gen_class add_locale_i certify_class; |
|
19038 | 450 |
|
451 |
end; (* local *) |
|
452 |
||
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
453 |
local |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
454 |
|
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
455 |
fun gen_add_defs_overloaded prep_att tap_def add_defs tyco raw_defs thy = |
19038 | 456 |
let |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
457 |
fun invent_name raw_t = |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
458 |
let |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
459 |
val t = tap_def thy raw_t; |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
460 |
val c = (fst o dest_Const o fst o strip_comb o fst o Logic.dest_equals) t; |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
461 |
in |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
462 |
Thm.def_name (NameSpace.base c ^ "_" ^ NameSpace.base tyco) |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
463 |
end; |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
464 |
fun prep_def (_, (("", a), t)) = |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
465 |
let |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
466 |
val n = invent_name t |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
467 |
in ((n, t), map (prep_att thy) a) end |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
468 |
| prep_def (_, ((n, a), t)) = |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
469 |
((n, t), map (prep_att thy) a); |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
470 |
in |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
471 |
thy |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
472 |
|> add_defs true (map prep_def raw_defs) |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
473 |
end; |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
474 |
|
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
475 |
val add_defs_overloaded = gen_add_defs_overloaded Attrib.attribute Sign.read_term PureThy.add_defs; |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
476 |
val add_defs_overloaded_i = gen_add_defs_overloaded (K I) (K I) PureThy.add_defs_i; |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
477 |
|
19136 | 478 |
fun gen_instance_arity prep_arity prep_att add_defs tap_def do_proof raw_arity (raw_name, raw_atts) raw_defs theory = |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
479 |
let |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
480 |
val pp = Sign.pp theory; |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
481 |
val arity as (tyco, asorts, sort) = prep_arity theory ((fn ((x, y), z) => (x, y, z)) raw_arity); |
19038 | 482 |
val ty_inst = Type (tyco, map2 (curry TVar o rpair 0) (Term.invent_names [] "'a" (length asorts)) asorts) |
19136 | 483 |
val name = case raw_name |
484 |
of "" => Thm.def_name ((space_implode "_" o map NameSpace.base) sort ^ "_" ^ NameSpace.base tyco) |
|
485 |
| _ => raw_name; |
|
486 |
val atts = map (prep_att theory) raw_atts; |
|
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
487 |
fun get_classes thy tyco sort = |
19038 | 488 |
let |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
489 |
fun get class classes = |
19213 | 490 |
if AList.defined (op =) ((the_instances thy) class) tyco |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
491 |
then classes |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
492 |
else classes |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
493 |
|> cons class |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
494 |
|> fold get (the_superclasses thy class) |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
495 |
in fold get sort [] end; |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
496 |
val classes = get_classes theory tyco sort; |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
497 |
val _ = if null classes then error ("already instantiated") else (); |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
498 |
fun get_consts class = |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
499 |
let |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
500 |
val data = the_class_data theory class; |
19038 | 501 |
val subst_ty = map_type_tfree (fn (var as (v, _)) => |
502 |
if #var data = v then ty_inst else TFree var) |
|
19280
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
503 |
in (map (apsnd subst_ty o snd) o #consts) data end; |
19468 | 504 |
val cs = (flat o map get_consts) classes; |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
505 |
fun get_remove_contraint c thy = |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
506 |
let |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
507 |
val ty = Sign.the_const_constraint thy c; |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
508 |
in |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
509 |
thy |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
510 |
|> Sign.add_const_constraint_i (c, NONE) |
19136 | 511 |
|> pair (c, Type.unvarifyT ty) |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
512 |
end; |
19253 | 513 |
fun check_defs0 thy raw_defs c_req = |
19038 | 514 |
let |
515 |
fun get_c raw_def = |
|
19253 | 516 |
(fst o Sign.cert_def pp o tap_def thy o snd) raw_def; |
19038 | 517 |
val c_given = map get_c raw_defs; |
19340 | 518 |
fun eq_c ((c1 : string, ty1), (c2, ty2)) = |
19038 | 519 |
let |
520 |
val ty1' = Type.varifyT ty1; |
|
521 |
val ty2' = Type.varifyT ty2; |
|
522 |
in |
|
523 |
c1 = c2 |
|
19253 | 524 |
andalso Sign.typ_instance thy (ty1', ty2') |
525 |
andalso Sign.typ_instance thy (ty2', ty1') |
|
19038 | 526 |
end; |
19300 | 527 |
val _ = case subtract eq_c c_req c_given |
19038 | 528 |
of [] => () |
529 |
| cs => error ("superfluous definition(s) given for " |
|
19253 | 530 |
^ (commas o map (fn (c, ty) => quote (c ^ "::" ^ Sign.string_of_typ thy ty))) cs); |
19340 | 531 |
(*val _ = case subtract eq_c c_given c_req |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
532 |
of [] => () |
19038 | 533 |
| cs => error ("no definition(s) given for " |
19340 | 534 |
^ (commas o map (fn (c, ty) => quote (c ^ "::" ^ Sign.string_of_typ thy ty))) cs);*) |
19253 | 535 |
in () end; |
536 |
fun check_defs1 raw_defs c_req thy = |
|
537 |
let |
|
538 |
val thy' = (Sign.add_arities_i [(tyco, asorts, sort)] o Theory.copy) thy |
|
539 |
in (check_defs0 thy' raw_defs c_req; thy) end; |
|
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
540 |
fun mangle_alldef_name tyco sort = |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
541 |
Thm.def_name ((space_implode "_" o map NameSpace.base) sort ^ "_" ^ NameSpace.base tyco); |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
542 |
fun note_all tyco sort thms thy = |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
543 |
thy |
19136 | 544 |
|> PureThy.note_thmss_i PureThy.internalK [((name, atts), [(thms, [])])] |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
545 |
|> snd; |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
546 |
fun after_qed cs thy = |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
547 |
thy |
19253 | 548 |
|> fold (fn class => |
549 |
add_inst_data (class, (tyco, |
|
19412 | 550 |
(map (operational_sort_of thy) asorts, Context.theory_name thy)))) sort |
551 |
|> fold Sign.add_const_constraint_i (map (apsnd SOME) cs); |
|
19038 | 552 |
in |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
553 |
theory |
19253 | 554 |
|> check_defs1 raw_defs cs |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
555 |
|> fold_map get_remove_contraint (map fst cs) |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
556 |
||>> add_defs tyco (map (pair NONE) raw_defs) |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
557 |
|-> (fn (cs, defnames) => note_all tyco sort defnames #> pair cs) |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
558 |
|-> (fn cs => do_proof (after_qed cs) arity) |
19038 | 559 |
end; |
560 |
||
19246 | 561 |
fun instance_arity' do_proof = gen_instance_arity Sign.read_arity Attrib.attribute add_defs_overloaded |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
562 |
(fn thy => fn t => (snd o read_axm thy) ("", t)) do_proof; |
19246 | 563 |
fun instance_arity_i' do_proof = gen_instance_arity Sign.cert_arity (K I) add_defs_overloaded_i |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
564 |
(K I) do_proof; |
19246 | 565 |
val setup_proof = axclass_instance_arity_i; |
566 |
fun tactic_proof tac after_qed arity = AxClass.prove_arity arity tac #> after_qed; |
|
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
567 |
|
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
568 |
in |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
569 |
|
19150 | 570 |
val instance_arity = instance_arity' setup_proof; |
571 |
val instance_arity_i = instance_arity_i' setup_proof; |
|
572 |
val prove_instance_arity = instance_arity_i' o tactic_proof; |
|
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
573 |
|
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
574 |
end; (* local *) |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
575 |
|
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
576 |
local |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
577 |
|
19136 | 578 |
fun fish_thms (name, expr) after_qed thy = |
579 |
let |
|
580 |
val _ = writeln ("sub " ^ name) |
|
581 |
val suplocales = (fn Locale.Merge es => map (fn Locale.Locale n => n) es) expr; |
|
582 |
val _ = writeln ("super " ^ commas suplocales) |
|
19280
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
583 |
fun get_c name = |
19136 | 584 |
(map (NameSpace.base o fst o fst) o Locale.parameters_of thy) name; |
585 |
fun get_a name = |
|
586 |
(map (NameSpace.base o fst o fst) o Locale.local_asms_of thy) name; |
|
587 |
fun get_t supname = |
|
588 |
map (NameSpace.append (NameSpace.append name ((space_implode "_" o get_c) supname)) o NameSpace.base) |
|
589 |
(get_a name); |
|
590 |
val names = map get_t suplocales; |
|
591 |
val _ = writeln ("fishing for " ^ (commas o map commas) names); |
|
592 |
in |
|
593 |
thy |
|
594 |
|> after_qed ((map o map) (Drule.standard o get_thm thy o Name) names) |
|
595 |
end; |
|
596 |
||
597 |
fun add_interpretation_in (after_qed : thm list list -> theory -> theory) (name, expr) thy = |
|
598 |
thy |
|
599 |
|> Locale.interpretation_in_locale (name, expr); |
|
600 |
||
601 |
fun prove_interpretation_in tac (after_qed : thm list list -> theory -> theory) (name, expr) thy = |
|
602 |
thy |
|
603 |
|> Locale.interpretation_in_locale (name, expr) |
|
604 |
|> Proof.global_terminal_proof (Method.Basic (fn _ => Method.SIMPLE_METHOD tac), NONE) |
|
605 |
|-> (fn _ => I); |
|
19038 | 606 |
|
19150 | 607 |
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
|
608 |
let |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
609 |
val class = prep_class theory raw_class; |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
610 |
val sort = prep_sort theory raw_sort; |
19136 | 611 |
val loc_name = (#name_locale o the_class_data theory) class; |
19468 | 612 |
val loc_expr = |
613 |
(Locale.Merge o map (Locale.Locale o #name_locale o the_class_data theory)) sort; |
|
19136 | 614 |
fun after_qed thmss thy = |
19468 | 615 |
(writeln "---"; (Pretty.writeln o Display.pretty_thms o flat) thmss; writeln "---"; fold (fn supclass => |
19412 | 616 |
AxClass.prove_classrel (class, supclass) |
19136 | 617 |
(ALLGOALS (K (intro_classes_tac [])) THEN |
19468 | 618 |
(ALLGOALS o resolve_tac o flat) thmss) |
19136 | 619 |
) sort thy) |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
620 |
in |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
621 |
theory |
19136 | 622 |
|> do_proof after_qed (loc_name, loc_expr) |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
623 |
end; |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
624 |
|
19150 | 625 |
fun instance_sort' do_proof = gen_instance_sort intern_class read_sort do_proof; |
626 |
fun instance_sort_i' do_proof = gen_instance_sort certify_class certify_sort do_proof; |
|
19136 | 627 |
val setup_proof = add_interpretation_in; |
628 |
val tactic_proof = prove_interpretation_in; |
|
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
629 |
|
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
630 |
in |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
631 |
|
19150 | 632 |
val instance_sort = instance_sort' setup_proof; |
633 |
val instance_sort_i = instance_sort_i' setup_proof; |
|
634 |
val prove_instance_sort = instance_sort_i' o tactic_proof; |
|
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
635 |
|
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
636 |
end; (* local *) |
19038 | 637 |
|
18168 | 638 |
(* extracting dictionary obligations from types *) |
639 |
||
640 |
type sortcontext = (string * sort) list; |
|
641 |
||
18335 | 642 |
fun extract_sortctxt thy ty = |
18702 | 643 |
(typ_tfrees o fst o Type.freeze_thaw_type) ty |
18884 | 644 |
|> map (apsnd (operational_sort_of thy)) |
18168 | 645 |
|> filter (not o null o snd); |
646 |
||
18884 | 647 |
datatype classlookup = Instance of (class * string) * classlookup list list |
19253 | 648 |
| Lookup of class list * (string * (int * int)) |
19213 | 649 |
|
650 |
fun pretty_lookup' (Instance ((class, tyco), lss)) = |
|
651 |
(Pretty.block o Pretty.breaks) ( |
|
652 |
Pretty.enum "," "{" "}" [Pretty.str class, Pretty.str tyco] |
|
653 |
:: map pretty_lookup lss |
|
654 |
) |
|
19253 | 655 |
| pretty_lookup' (Lookup (classes, (v, (i, j)))) = |
656 |
Pretty.enum " <" "[" "]" (map Pretty.str classes @ [Pretty.str (v ^ "!" ^ string_of_int i ^ "/" ^ string_of_int j)]) |
|
19213 | 657 |
and pretty_lookup ls = (Pretty.enum "," "(" ")" o map pretty_lookup') ls; |
18168 | 658 |
|
18864 | 659 |
fun extract_lookup thy sortctxt raw_typ_def raw_typ_use = |
18168 | 660 |
let |
661 |
val typ_def = Type.varifyT raw_typ_def; |
|
662 |
val typ_use = Type.varifyT raw_typ_use; |
|
663 |
val match_tab = Sign.typ_match thy (typ_def, typ_use) Vartab.empty; |
|
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
664 |
fun tab_lookup vname = (the o Vartab.lookup match_tab) (vname, 0); |
18168 | 665 |
fun mk_class_deriv thy subclasses superclass = |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
666 |
let |
19233
77ca20b0ed77
renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
haftmann
parents:
19213
diff
changeset
|
667 |
val (i, (subclass::deriv)) = (the oo get_index) (fn subclass => |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
668 |
get_superclass_derivation thy (subclass, superclass) |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
669 |
) subclasses; |
19253 | 670 |
in (rev deriv, (i, length subclasses)) end; |
19213 | 671 |
fun mk_lookup (sort_def, (Type (tyco, tys))) = |
672 |
map (fn class => Instance ((class, tyco), |
|
673 |
map2 (curry mk_lookup) |
|
19468 | 674 |
(map (operational_sort_of thy) (Sign.arity_sorts thy tyco [class])) |
19213 | 675 |
tys) |
676 |
) sort_def |
|
18168 | 677 |
| mk_lookup (sort_def, TVar ((vname, _), sort_use)) = |
678 |
let |
|
679 |
fun mk_look class = |
|
18884 | 680 |
let val (deriv, classindex) = mk_class_deriv thy (operational_sort_of thy sort_use) class |
18168 | 681 |
in Lookup (deriv, (vname, classindex)) end; |
682 |
in map mk_look sort_def end; |
|
18864 | 683 |
in |
19280
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
684 |
sortctxt |
18864 | 685 |
|> map (tab_lookup o fst) |
18884 | 686 |
|> map (apfst (operational_sort_of thy)) |
18864 | 687 |
|> filter (not o null o fst) |
688 |
|> map mk_lookup |
|
689 |
end; |
|
690 |
||
18884 | 691 |
fun extract_classlookup thy (c, raw_typ_use) = |
18864 | 692 |
let |
693 |
val raw_typ_def = Sign.the_const_constraint thy c; |
|
694 |
val typ_def = Type.varifyT raw_typ_def; |
|
18702 | 695 |
fun reorder_sortctxt ctxt = |
696 |
case lookup_const_class thy c |
|
697 |
of NONE => ctxt |
|
698 |
| SOME class => |
|
699 |
let |
|
19038 | 700 |
val data = the_class_data thy class; |
19280
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
701 |
val sign = (Type.varifyT o the o AList.lookup (op =) ((map snd o #consts) data)) c; |
18702 | 702 |
val match_tab = Sign.typ_match thy (sign, typ_def) Vartab.empty; |
703 |
val v : string = case Vartab.lookup match_tab (#var data, 0) |
|
704 |
of SOME (_, TVar ((v, _), _)) => v; |
|
705 |
in |
|
706 |
(v, (the o AList.lookup (op =) ctxt) v) :: AList.delete (op =) v ctxt |
|
707 |
end; |
|
18168 | 708 |
in |
18864 | 709 |
extract_lookup thy |
710 |
(reorder_sortctxt (extract_sortctxt thy ((fst o Type.freeze_thaw_type) raw_typ_def))) |
|
711 |
raw_typ_def raw_typ_use |
|
18168 | 712 |
end; |
713 |
||
18884 | 714 |
fun extract_classlookup_inst thy (class, tyco) supclass = |
18864 | 715 |
let |
716 |
fun mk_typ class = Type (tyco, (map TFree o fst o the_inst_sign thy) (class, tyco)) |
|
19136 | 717 |
val typ_def = mk_typ supclass; |
718 |
val typ_use = mk_typ class; |
|
18864 | 719 |
in |
720 |
extract_lookup thy (extract_sortctxt thy typ_def) typ_def typ_use |
|
721 |
end; |
|
18168 | 722 |
|
19253 | 723 |
fun extract_classlookup_member thy (ty_decl, ty_use) = |
724 |
extract_lookup thy (extract_sortctxt thy ty_decl) ty_decl ty_use; |
|
18884 | 725 |
|
18515 | 726 |
(* toplevel interface *) |
727 |
||
728 |
local |
|
729 |
||
730 |
structure P = OuterParse |
|
731 |
and K = OuterKeyword |
|
732 |
||
733 |
in |
|
734 |
||
18849
05a16861d3a5
added three times overloaded Isar instance command
haftmann
parents:
18829
diff
changeset
|
735 |
val (classK, instanceK) = ("class", "instance") |
18515 | 736 |
|
19253 | 737 |
val use_cp_instance = ref false; |
738 |
||
19136 | 739 |
fun wrap_add_instance_subclass (class, sort) thy = |
740 |
case Sign.read_sort thy sort |
|
741 |
of [class'] => |
|
19253 | 742 |
if ! use_cp_instance |
743 |
andalso (is_some o lookup_class_data thy o Sign.intern_class thy) class |
|
19136 | 744 |
andalso (is_some o lookup_class_data thy o Sign.intern_class thy) class' |
745 |
then |
|
19150 | 746 |
instance_sort (class, sort) thy |
19136 | 747 |
else |
19246 | 748 |
axclass_instance_subclass (class, sort) thy |
19150 | 749 |
| _ => instance_sort (class, sort) thy; |
19136 | 750 |
|
18911 | 751 |
val parse_inst = |
19136 | 752 |
(Scan.optional (P.$$$ "(" |-- P.!!! (P.list1 P.sort --| P.$$$ ")")) [] -- P.xname --| P.$$$ "::" -- P.sort) |
753 |
>> (fn ((asorts, tyco), sort) => ((tyco, asorts), sort)) |
|
754 |
|| (P.xname --| P.$$$ "::" -- P.!!! P.arity) |
|
755 |
>> (fn (tyco, (asorts, sort)) => ((tyco, asorts), sort)); |
|
18911 | 756 |
|
19038 | 757 |
val locale_val = |
758 |
(P.locale_expr -- |
|
759 |
Scan.optional (P.$$$ "+" |-- P.!!! (Scan.repeat1 P.context_element)) [] || |
|
760 |
Scan.repeat1 P.context_element >> pair Locale.empty); |
|
761 |
||
19280
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
762 |
val class_subP = P.name -- Scan.repeat (P.$$$ "+" |-- P.name) >> (op ::); |
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
763 |
val class_bodyP = P.!!! (Scan.repeat1 P.context_element); |
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
764 |
|
18515 | 765 |
val classP = |
18849
05a16861d3a5
added three times overloaded Isar instance command
haftmann
parents:
18829
diff
changeset
|
766 |
OuterSyntax.command classK "operational type classes" K.thy_decl ( |
05a16861d3a5
added three times overloaded Isar instance command
haftmann
parents:
18829
diff
changeset
|
767 |
P.name --| P.$$$ "=" |
19280
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
768 |
-- ( |
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
769 |
class_subP --| P.$$$ "+" -- class_bodyP |
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
770 |
|| class_subP >> rpair [] |
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
771 |
|| class_bodyP >> pair [] |
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
772 |
) >> (Toplevel.theory_context |
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
773 |
o (fn (bname, (supclasses, elems)) => class bname supclasses elems))); |
18515 | 774 |
|
18575 | 775 |
val instanceP = |
18849
05a16861d3a5
added three times overloaded Isar instance command
haftmann
parents:
18829
diff
changeset
|
776 |
OuterSyntax.command instanceK "prove type arity or subclass relation" K.thy_goal (( |
19136 | 777 |
P.xname -- ((P.$$$ "\\<subseteq>" || P.$$$ "<") |-- P.!!! P.xname) >> wrap_add_instance_subclass |
778 |
|| P.opt_thm_name ":" -- (parse_inst -- Scan.repeat (P.opt_thm_name ":" -- P.prop)) |
|
19246 | 779 |
>> (fn (("", []), (((tyco, asorts), sort), [])) => axclass_instance_arity I (tyco, asorts, sort) |
19150 | 780 |
| (natts, (inst, defs)) => instance_arity inst natts defs) |
18849
05a16861d3a5
added three times overloaded Isar instance command
haftmann
parents:
18829
diff
changeset
|
781 |
) >> (Toplevel.print oo Toplevel.theory_to_proof)); |
18575 | 782 |
|
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
783 |
val _ = OuterSyntax.add_parsers [classP, instanceP]; |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
784 |
|
18515 | 785 |
end; (* local *) |
786 |
||
18168 | 787 |
end; (* struct *) |