| author | wenzelm |
| Fri, 16 Jan 2009 21:24:33 +0100 | |
| changeset 29520 | 7402322256b0 |
| parent 29509 | 1ff0f3f08a7b |
| child 29526 | 0b32c8b84d3e |
| permissions | -rw-r--r-- |
| 29358 | 1 |
(* Title: Pure/Isar/class_target.ML |
| 24218 | 2 |
Author: Florian Haftmann, TU Muenchen |
3 |
||
| 29358 | 4 |
Type classes derived from primitive axclasses and locales - mechanisms. |
| 24218 | 5 |
*) |
6 |
||
| 29358 | 7 |
signature CLASS_TARGET = |
| 24218 | 8 |
sig |
| 25462 | 9 |
(*classes*) |
| 29358 | 10 |
val register: class -> class list -> ((string * typ) * (string * typ)) list |
| 29439 | 11 |
-> sort -> term list -> morphism |
| 29358 | 12 |
-> thm option -> thm option -> thm -> theory -> theory |
|
29509
1ff0f3f08a7b
migrated class package to new locale implementation
haftmann
parents:
29439
diff
changeset
|
13 |
val register_subclass: class * class -> thm option |
|
1ff0f3f08a7b
migrated class package to new locale implementation
haftmann
parents:
29439
diff
changeset
|
14 |
-> theory -> theory |
| 25485 | 15 |
|
| 29358 | 16 |
val begin: class list -> sort -> Proof.context -> Proof.context |
| 25311 | 17 |
val init: class -> theory -> Proof.context |
| 28017 | 18 |
val declare: class -> Properties.T |
|
27690
24738db98d34
some steps towards explicit class target for canonical interpretation
haftmann
parents:
27684
diff
changeset
|
19 |
-> (string * mixfix) * term -> theory -> theory |
| 28017 | 20 |
val abbrev: class -> Syntax.mode -> Properties.T |
| 25104 | 21 |
-> (string * mixfix) * term -> theory -> theory |
| 25083 | 22 |
val refresh_syntax: class -> Proof.context -> Proof.context |
| 25485 | 23 |
|
| 24589 | 24 |
val intro_classes_tac: thm list -> tactic |
| 26470 | 25 |
val default_intro_tac: Proof.context -> thm list -> tactic |
| 25485 | 26 |
|
27 |
val class_prefix: string -> string |
|
28 |
val is_class: theory -> class -> bool |
|
| 26518 | 29 |
val these_params: theory -> sort -> (string * (class * (string * typ))) list |
| 29358 | 30 |
val these_defs: theory -> sort -> thm list |
|
29509
1ff0f3f08a7b
migrated class package to new locale implementation
haftmann
parents:
29439
diff
changeset
|
31 |
val eq_morph: theory -> thm list -> morphism |
| 29358 | 32 |
val base_sort: theory -> class -> sort |
33 |
val rules: theory -> class -> thm option * thm |
|
| 24589 | 34 |
val print_classes: theory -> unit |
|
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
35 |
|
| 25462 | 36 |
(*instances*) |
| 26247 | 37 |
val init_instantiation: string list * (string * sort) list * sort |
38 |
-> theory -> local_theory |
|
39 |
val instantiation_instance: (local_theory -> local_theory) |
|
40 |
-> local_theory -> Proof.state |
|
41 |
val prove_instantiation_instance: (Proof.context -> tactic) |
|
42 |
-> local_theory -> local_theory |
|
| 28666 | 43 |
val prove_instantiation_exit: (Proof.context -> tactic) |
44 |
-> local_theory -> theory |
|
45 |
val prove_instantiation_exit_result: (morphism -> 'a -> 'b) |
|
46 |
-> (Proof.context -> 'b -> tactic) -> 'a -> local_theory -> 'b * theory |
|
| 25485 | 47 |
val conclude_instantiation: local_theory -> local_theory |
| 25603 | 48 |
val instantiation_param: local_theory -> string -> string option |
| 25485 | 49 |
val confirm_declaration: string -> local_theory -> local_theory |
| 25603 | 50 |
val pretty_instantiation: local_theory -> Pretty.T |
| 26259 | 51 |
val type_name: string -> string |
| 25485 | 52 |
|
| 25462 | 53 |
(*old axclass layer*) |
54 |
val axclass_cmd: bstring * xstring list |
|
|
28084
a05ca48ef263
type Attrib.binding abbreviates Name.binding without attributes;
wenzelm
parents:
28083
diff
changeset
|
55 |
-> (Attrib.binding * string list) list |
| 25462 | 56 |
-> theory -> class * theory |
57 |
val classrel_cmd: xstring * xstring -> theory -> Proof.state |
|
58 |
||
59 |
(*old instance layer*) |
|
| 25536 | 60 |
val instance_arity: (theory -> theory) -> arity -> theory -> Proof.state |
61 |
val instance_arity_cmd: bstring * xstring list * xstring -> theory -> Proof.state |
|
| 24218 | 62 |
end; |
63 |
||
| 29358 | 64 |
structure Class_Target : CLASS_TARGET = |
| 24218 | 65 |
struct |
66 |
||
| 25485 | 67 |
(** primitive axclass and instance commands **) |
| 24589 | 68 |
|
| 24218 | 69 |
fun axclass_cmd (class, raw_superclasses) raw_specs thy = |
70 |
let |
|
71 |
val ctxt = ProofContext.init thy; |
|
72 |
val superclasses = map (Sign.read_class thy) raw_superclasses; |
|
| 24589 | 73 |
val name_atts = map ((apsnd o map) (Attrib.attribute thy) o fst) |
74 |
raw_specs; |
|
75 |
val axiomss = ProofContext.read_propp (ctxt, map (map (rpair []) o snd) |
|
76 |
raw_specs) |
|
| 24218 | 77 |
|> snd |
78 |
|> (map o map) fst; |
|
| 24589 | 79 |
in |
80 |
AxClass.define_class (class, superclasses) [] |
|
81 |
(name_atts ~~ axiomss) thy |
|
82 |
end; |
|
| 24218 | 83 |
|
84 |
local |
|
85 |
||
86 |
fun gen_instance mk_prop add_thm after_qed insts thy = |
|
87 |
let |
|
88 |
fun after_qed' results = |
|
89 |
ProofContext.theory ((fold o fold) add_thm results #> after_qed); |
|
90 |
in |
|
91 |
thy |
|
92 |
|> ProofContext.init |
|
| 24589 | 93 |
|> Proof.theorem_i NONE after_qed' ((map (fn t => [(t, [])]) |
| 25536 | 94 |
o mk_prop thy) insts) |
| 24218 | 95 |
end; |
96 |
||
97 |
in |
|
98 |
||
| 24589 | 99 |
val instance_arity = |
| 24218 | 100 |
gen_instance (Logic.mk_arities oo Sign.cert_arity) AxClass.add_arity; |
| 25502 | 101 |
val instance_arity_cmd = |
102 |
gen_instance (Logic.mk_arities oo Sign.read_arity) AxClass.add_arity I; |
|
| 24589 | 103 |
val classrel = |
| 25536 | 104 |
gen_instance (single oo (Logic.mk_classrel oo AxClass.cert_classrel)) AxClass.add_classrel I; |
| 24589 | 105 |
val classrel_cmd = |
| 25536 | 106 |
gen_instance (single oo (Logic.mk_classrel oo AxClass.read_classrel)) AxClass.add_classrel I; |
| 24218 | 107 |
|
108 |
end; (*local*) |
|
109 |
||
110 |
||
| 24589 | 111 |
(** class data **) |
| 24218 | 112 |
|
113 |
datatype class_data = ClassData of {
|
|
|
28715
238f9966c80e
class morphism stemming from locale interpretation
haftmann
parents:
28674
diff
changeset
|
114 |
|
|
238f9966c80e
class morphism stemming from locale interpretation
haftmann
parents:
28674
diff
changeset
|
115 |
(* static part *) |
| 24218 | 116 |
consts: (string * string) list |
| 24836 | 117 |
(*locale parameter ~> constant name*), |
| 25062 | 118 |
base_sort: sort, |
|
28715
238f9966c80e
class morphism stemming from locale interpretation
haftmann
parents:
28674
diff
changeset
|
119 |
inst: term list |
| 25083 | 120 |
(*canonical interpretation*), |
| 29439 | 121 |
base_morph: Morphism.morphism |
122 |
(*static part of canonical morphism*), |
|
| 25618 | 123 |
assm_intro: thm option, |
124 |
of_class: thm, |
|
125 |
axiom: thm option, |
|
|
28715
238f9966c80e
class morphism stemming from locale interpretation
haftmann
parents:
28674
diff
changeset
|
126 |
|
|
238f9966c80e
class morphism stemming from locale interpretation
haftmann
parents:
28674
diff
changeset
|
127 |
(* dynamic part *) |
| 24657 | 128 |
defs: thm list, |
|
25368
f12613fda79d
proper implementation of check phase; non-qualified names for class operations
haftmann
parents:
25344
diff
changeset
|
129 |
operations: (string * (class * (typ * term))) list |
|
28715
238f9966c80e
class morphism stemming from locale interpretation
haftmann
parents:
28674
diff
changeset
|
130 |
|
| 24657 | 131 |
}; |
| 24218 | 132 |
|
| 24657 | 133 |
fun rep_class_data (ClassData d) = d; |
| 29439 | 134 |
fun mk_class_data ((consts, base_sort, inst, base_morph, assm_intro, of_class, axiom), |
|
25368
f12613fda79d
proper implementation of check phase; non-qualified names for class operations
haftmann
parents:
25344
diff
changeset
|
135 |
(defs, operations)) = |
| 25062 | 136 |
ClassData { consts = consts, base_sort = base_sort, inst = inst,
|
| 29439 | 137 |
base_morph = base_morph, assm_intro = assm_intro, of_class = of_class, axiom = axiom, |
| 25618 | 138 |
defs = defs, operations = operations }; |
| 29439 | 139 |
fun map_class_data f (ClassData { consts, base_sort, inst, base_morph, assm_intro,
|
|
28715
238f9966c80e
class morphism stemming from locale interpretation
haftmann
parents:
28674
diff
changeset
|
140 |
of_class, axiom, defs, operations }) = |
| 29439 | 141 |
mk_class_data (f ((consts, base_sort, inst, base_morph, assm_intro, of_class, axiom), |
|
25368
f12613fda79d
proper implementation of check phase; non-qualified names for class operations
haftmann
parents:
25344
diff
changeset
|
142 |
(defs, operations))); |
| 25038 | 143 |
fun merge_class_data _ (ClassData { consts = consts,
|
| 29439 | 144 |
base_sort = base_sort, inst = inst, base_morph = base_morph, assm_intro = assm_intro, |
| 25618 | 145 |
of_class = of_class, axiom = axiom, defs = defs1, operations = operations1 }, |
| 29439 | 146 |
ClassData { consts = _, base_sort = _, inst = _, base_morph = _, assm_intro = _,
|
| 25618 | 147 |
of_class = _, axiom = _, defs = defs2, operations = operations2 }) = |
| 29439 | 148 |
mk_class_data ((consts, base_sort, inst, base_morph, assm_intro, of_class, axiom), |
|
24914
95cda5dd58d5
added proper subclass concept; improved class target
haftmann
parents:
24901
diff
changeset
|
149 |
(Thm.merge_thms (defs1, defs2), |
|
25368
f12613fda79d
proper implementation of check phase; non-qualified names for class operations
haftmann
parents:
25344
diff
changeset
|
150 |
AList.merge (op =) (K true) (operations1, operations2))); |
| 24218 | 151 |
|
152 |
structure ClassData = TheoryDataFun |
|
153 |
( |
|
| 25038 | 154 |
type T = class_data Graph.T |
155 |
val empty = Graph.empty; |
|
| 24218 | 156 |
val copy = I; |
157 |
val extend = I; |
|
| 25038 | 158 |
fun merge _ = Graph.join merge_class_data; |
| 24218 | 159 |
); |
160 |
||
161 |
||
162 |
(* queries *) |
|
163 |
||
| 25038 | 164 |
val lookup_class_data = Option.map rep_class_data oo try o Graph.get_node o ClassData.get; |
| 24218 | 165 |
|
| 24589 | 166 |
fun the_class_data thy class = case lookup_class_data thy class |
| 25020 | 167 |
of NONE => error ("Undeclared class " ^ quote class)
|
| 24589 | 168 |
| SOME data => data; |
| 24218 | 169 |
|
| 25038 | 170 |
val is_class = is_some oo lookup_class_data; |
171 |
||
172 |
val ancestry = Graph.all_succs o ClassData.get; |
|
| 24218 | 173 |
|
|
29509
1ff0f3f08a7b
migrated class package to new locale implementation
haftmann
parents:
29439
diff
changeset
|
174 |
val heritage = Graph.all_preds o ClassData.get; |
|
1ff0f3f08a7b
migrated class package to new locale implementation
haftmann
parents:
29439
diff
changeset
|
175 |
|
|
28715
238f9966c80e
class morphism stemming from locale interpretation
haftmann
parents:
28674
diff
changeset
|
176 |
fun the_inst thy = #inst o the_class_data thy; |
|
238f9966c80e
class morphism stemming from locale interpretation
haftmann
parents:
28674
diff
changeset
|
177 |
|
| 25002 | 178 |
fun these_params thy = |
| 24218 | 179 |
let |
180 |
fun params class = |
|
181 |
let |
|
|
24930
cc2e0e8c81af
renamed AxClass.get_definition to AxClass.get_info (again);
wenzelm
parents:
24920
diff
changeset
|
182 |
val const_typs = (#params o AxClass.get_info thy) class; |
| 24657 | 183 |
val const_names = (#consts o the_class_data thy) class; |
| 24218 | 184 |
in |
| 26518 | 185 |
(map o apsnd) |
186 |
(fn c => (class, (c, (the o AList.lookup (op =) const_typs) c))) const_names |
|
| 24218 | 187 |
end; |
188 |
in maps params o ancestry thy end; |
|
189 |
||
| 25618 | 190 |
fun these_assm_intros thy = |
191 |
Graph.fold (fn (_, (data, _)) => fold (insert Thm.eq_thm) |
|
192 |
((the_list o #assm_intro o rep_class_data) data)) (ClassData.get thy) []; |
|
| 24218 | 193 |
|
|
28715
238f9966c80e
class morphism stemming from locale interpretation
haftmann
parents:
28674
diff
changeset
|
194 |
fun these_defs thy = maps (these o Option.map #defs o lookup_class_data thy) o ancestry thy; |
|
238f9966c80e
class morphism stemming from locale interpretation
haftmann
parents:
28674
diff
changeset
|
195 |
|
| 24836 | 196 |
fun these_operations thy = |
197 |
maps (#operations o the_class_data thy) o ancestry thy; |
|
| 24657 | 198 |
|
| 29358 | 199 |
val base_sort = #base_sort oo the_class_data; |
200 |
||
201 |
fun rules thy class = let |
|
202 |
val { axiom, of_class, ... } = the_class_data thy class
|
|
203 |
in (axiom, of_class) end; |
|
204 |
||
| 29439 | 205 |
val class_prefix = Logic.const_of_class o Sign.base_name; |
206 |
||
207 |
fun class_binding_morph class = |
|
208 |
Binding.map_prefix (K (Binding.add_prefix false (class_prefix class))); |
|
209 |
||
|
29509
1ff0f3f08a7b
migrated class package to new locale implementation
haftmann
parents:
29439
diff
changeset
|
210 |
fun eq_morph thy thms = (*FIXME how general is this?*) |
|
1ff0f3f08a7b
migrated class package to new locale implementation
haftmann
parents:
29439
diff
changeset
|
211 |
Morphism.term_morphism (MetaSimplifier.rewrite_term thy thms []) |
|
1ff0f3f08a7b
migrated class package to new locale implementation
haftmann
parents:
29439
diff
changeset
|
212 |
$> Morphism.thm_morphism (MetaSimplifier.rewrite_rule thms); |
|
1ff0f3f08a7b
migrated class package to new locale implementation
haftmann
parents:
29439
diff
changeset
|
213 |
|
|
28715
238f9966c80e
class morphism stemming from locale interpretation
haftmann
parents:
28674
diff
changeset
|
214 |
fun morphism thy class = |
|
238f9966c80e
class morphism stemming from locale interpretation
haftmann
parents:
28674
diff
changeset
|
215 |
let |
| 29439 | 216 |
val { base_morph, defs, ... } = the_class_data thy class;
|
|
29509
1ff0f3f08a7b
migrated class package to new locale implementation
haftmann
parents:
29439
diff
changeset
|
217 |
in base_morph $> eq_morph thy defs end; |
|
28715
238f9966c80e
class morphism stemming from locale interpretation
haftmann
parents:
28674
diff
changeset
|
218 |
|
| 24218 | 219 |
fun print_classes thy = |
220 |
let |
|
| 24920 | 221 |
val ctxt = ProofContext.init thy; |
| 24218 | 222 |
val algebra = Sign.classes_of thy; |
223 |
val arities = |
|
224 |
Symtab.empty |
|
225 |
|> Symtab.fold (fn (tyco, arities) => fold (fn (class, _) => |
|
226 |
Symtab.map_default (class, []) (insert (op =) tyco)) arities) |
|
227 |
((#arities o Sorts.rep_algebra) algebra); |
|
228 |
val the_arities = these o Symtab.lookup arities; |
|
229 |
fun mk_arity class tyco = |
|
230 |
let |
|
231 |
val Ss = Sorts.mg_domain algebra tyco [class]; |
|
| 24920 | 232 |
in Syntax.pretty_arity ctxt (tyco, Ss, [class]) end; |
| 24218 | 233 |
fun mk_param (c, ty) = Pretty.str (Sign.extern_const thy c ^ " :: " |
| 24920 | 234 |
^ setmp show_sorts false (Syntax.string_of_typ ctxt o Type.strip_sorts) ty); |
| 24218 | 235 |
fun mk_entry class = (Pretty.block o Pretty.fbreaks o map_filter I) [ |
| 25062 | 236 |
(SOME o Pretty.str) ("class " ^ Sign.extern_class thy class ^ ":"),
|
| 24218 | 237 |
(SOME o Pretty.block) [Pretty.str "supersort: ", |
| 24920 | 238 |
(Syntax.pretty_sort ctxt o Sign.minimize_sort thy o Sign.super_classes thy) class], |
| 25062 | 239 |
if is_class thy class then (SOME o Pretty.str) |
|
29509
1ff0f3f08a7b
migrated class package to new locale implementation
haftmann
parents:
29439
diff
changeset
|
240 |
("locale: " ^ Locale.extern thy class) else NONE,
|
| 25062 | 241 |
((fn [] => NONE | ps => (SOME o Pretty.block o Pretty.fbreaks) |
242 |
(Pretty.str "parameters:" :: ps)) o map mk_param |
|
|
24930
cc2e0e8c81af
renamed AxClass.get_definition to AxClass.get_info (again);
wenzelm
parents:
24920
diff
changeset
|
243 |
o these o Option.map #params o try (AxClass.get_info thy)) class, |
| 24218 | 244 |
(SOME o Pretty.block o Pretty.breaks) [ |
245 |
Pretty.str "instances:", |
|
246 |
Pretty.list "" "" (map (mk_arity class) (the_arities class)) |
|
247 |
] |
|
248 |
] |
|
249 |
in |
|
| 24589 | 250 |
(Pretty.writeln o Pretty.chunks o separate (Pretty.str "") |
251 |
o map mk_entry o Sorts.all_classes) algebra |
|
| 24218 | 252 |
end; |
253 |
||
254 |
||
255 |
(* updaters *) |
|
256 |
||
| 29439 | 257 |
fun register class superclasses params base_sort inst morph |
| 29358 | 258 |
axiom assm_intro of_class thy = |
| 25002 | 259 |
let |
|
25368
f12613fda79d
proper implementation of check phase; non-qualified names for class operations
haftmann
parents:
25344
diff
changeset
|
260 |
val operations = map (fn (v_ty as (_, ty), (c, _)) => |
| 25683 | 261 |
(c, (class, (ty, Free v_ty)))) params; |
| 25038 | 262 |
val add_class = Graph.new_node (class, |
| 25683 | 263 |
mk_class_data (((map o pairself) fst params, base_sort, |
| 29439 | 264 |
inst, morph, assm_intro, of_class, axiom), ([], operations))) |
| 25002 | 265 |
#> fold (curry Graph.add_edge class) superclasses; |
| 25618 | 266 |
in ClassData.map add_class thy end; |
| 24218 | 267 |
|
|
25368
f12613fda79d
proper implementation of check phase; non-qualified names for class operations
haftmann
parents:
25344
diff
changeset
|
268 |
fun register_operation class (c, (t, some_def)) thy = |
| 25062 | 269 |
let |
| 29358 | 270 |
val base_sort = base_sort thy class; |
| 29439 | 271 |
val prep_typ = map_type_tfree |
272 |
(fn (v, sort) => if Name.aT = v |
|
273 |
then TFree (v, base_sort) else TVar ((v, 0), sort)); |
|
|
25368
f12613fda79d
proper implementation of check phase; non-qualified names for class operations
haftmann
parents:
25344
diff
changeset
|
274 |
val t' = map_types prep_typ t; |
|
f12613fda79d
proper implementation of check phase; non-qualified names for class operations
haftmann
parents:
25344
diff
changeset
|
275 |
val ty' = Term.fastype_of t'; |
| 25062 | 276 |
in |
277 |
thy |
|
278 |
|> (ClassData.map o Graph.map_node class o map_class_data o apsnd) |
|
|
25368
f12613fda79d
proper implementation of check phase; non-qualified names for class operations
haftmann
parents:
25344
diff
changeset
|
279 |
(fn (defs, operations) => |
| 25096 | 280 |
(fold cons (the_list some_def) defs, |
|
25368
f12613fda79d
proper implementation of check phase; non-qualified names for class operations
haftmann
parents:
25344
diff
changeset
|
281 |
(c, (class, (ty', t'))) :: operations)) |
| 25062 | 282 |
end; |
| 24218 | 283 |
|
| 24589 | 284 |
|
| 29358 | 285 |
(** tactics and methods **) |
| 24589 | 286 |
|
|
29509
1ff0f3f08a7b
migrated class package to new locale implementation
haftmann
parents:
29439
diff
changeset
|
287 |
fun register_subclass (sub, sup) thms thy = |
| 25618 | 288 |
let |
| 29358 | 289 |
val of_class = (snd o rules thy) sup; |
|
29509
1ff0f3f08a7b
migrated class package to new locale implementation
haftmann
parents:
29439
diff
changeset
|
290 |
val intro = case thms |
|
27684
f45fd1159a4b
subclass now also works for subclasses with empty specificaton
haftmann
parents:
27281
diff
changeset
|
291 |
of SOME thm => Drule.standard' (of_class OF [Drule.standard' thm]) |
|
f45fd1159a4b
subclass now also works for subclasses with empty specificaton
haftmann
parents:
27281
diff
changeset
|
292 |
| NONE => Thm.instantiate ([pairself (Thm.ctyp_of thy o TVar o pair (Name.aT, 0)) |
|
f45fd1159a4b
subclass now also works for subclasses with empty specificaton
haftmann
parents:
27281
diff
changeset
|
293 |
([], [sub])], []) of_class; |
| 29358 | 294 |
val classrel = (intro OF (the_list o fst o rules thy) sub) |
|
27684
f45fd1159a4b
subclass now also works for subclasses with empty specificaton
haftmann
parents:
27281
diff
changeset
|
295 |
|> Thm.close_derivation; |
|
29509
1ff0f3f08a7b
migrated class package to new locale implementation
haftmann
parents:
29439
diff
changeset
|
296 |
(*FIXME generic amend operation for classes*) |
|
1ff0f3f08a7b
migrated class package to new locale implementation
haftmann
parents:
29439
diff
changeset
|
297 |
val amendments = map (fn class => (class, morphism thy class)) |
|
1ff0f3f08a7b
migrated class package to new locale implementation
haftmann
parents:
29439
diff
changeset
|
298 |
(heritage thy [sub]); |
|
1ff0f3f08a7b
migrated class package to new locale implementation
haftmann
parents:
29439
diff
changeset
|
299 |
val diff_sort = Sign.complete_sort thy [sup] |
|
1ff0f3f08a7b
migrated class package to new locale implementation
haftmann
parents:
29439
diff
changeset
|
300 |
|> subtract (op =) (Sign.complete_sort thy [sub]); |
|
1ff0f3f08a7b
migrated class package to new locale implementation
haftmann
parents:
29439
diff
changeset
|
301 |
val defs_morph = eq_morph thy (these_defs thy diff_sort); |
| 25618 | 302 |
in |
303 |
thy |
|
304 |
|> AxClass.add_classrel classrel |
|
305 |
|> ClassData.map (Graph.add_edge (sub, sup)) |
|
|
29509
1ff0f3f08a7b
migrated class package to new locale implementation
haftmann
parents:
29439
diff
changeset
|
306 |
|> fold (Locale.amend_registration defs_morph) amendments |
| 25618 | 307 |
end; |
308 |
||
| 24218 | 309 |
fun intro_classes_tac facts st = |
310 |
let |
|
311 |
val thy = Thm.theory_of_thm st; |
|
312 |
val classes = Sign.all_classes thy; |
|
313 |
val class_trivs = map (Thm.class_triv thy) classes; |
|
| 25618 | 314 |
val class_intros = map_filter (try (#intro o AxClass.get_info thy)) classes; |
315 |
val assm_intros = these_assm_intros thy; |
|
| 24218 | 316 |
in |
| 25618 | 317 |
Method.intros_tac (class_trivs @ class_intros @ assm_intros) facts st |
| 24218 | 318 |
end; |
319 |
||
| 26470 | 320 |
fun default_intro_tac ctxt [] = |
| 29362 | 321 |
intro_classes_tac [] ORELSE Old_Locale.intro_locales_tac true ctxt [] ORELSE |
|
29029
1945e0185097
NewLocale.intro_locales_tac added to Class.default_intro_tac.
ballarin
parents:
29006
diff
changeset
|
322 |
Locale.intro_locales_tac true ctxt [] |
| 26470 | 323 |
| default_intro_tac _ _ = no_tac; |
| 24218 | 324 |
|
325 |
fun default_tac rules ctxt facts = |
|
326 |
HEADGOAL (Method.some_rule_tac rules ctxt facts) ORELSE |
|
| 26470 | 327 |
default_intro_tac ctxt facts; |
| 24218 | 328 |
|
| 26463 | 329 |
val _ = Context.>> (Context.map_theory |
330 |
(Method.add_methods |
|
331 |
[("intro_classes", Method.no_args (Method.METHOD intro_classes_tac),
|
|
332 |
"back-chain introduction rules of classes"), |
|
333 |
("default", Method.thms_ctxt_args (Method.METHOD oo default_tac),
|
|
334 |
"apply some intro/elim rule")])); |
|
335 |
||
| 24218 | 336 |
|
| 24589 | 337 |
(** classes and class target **) |
| 24218 | 338 |
|
| 25002 | 339 |
(* class context syntax *) |
| 24748 | 340 |
|
| 26238 | 341 |
fun synchronize_class_syntax sups base_sort ctxt = |
|
24914
95cda5dd58d5
added proper subclass concept; improved class target
haftmann
parents:
24901
diff
changeset
|
342 |
let |
|
25344
00c2179db769
synchronize_syntax: improved declare_const (still inactive);
wenzelm
parents:
25326
diff
changeset
|
343 |
val thy = ProofContext.theory_of ctxt; |
| 26596 | 344 |
val algebra = Sign.classes_of thy; |
| 25083 | 345 |
val operations = these_operations thy sups; |
| 26518 | 346 |
fun subst_class_typ sort = map_type_tfree (K (TVar ((Name.aT, 0), sort))); |
347 |
val primary_constraints = |
|
|
25368
f12613fda79d
proper implementation of check phase; non-qualified names for class operations
haftmann
parents:
25344
diff
changeset
|
348 |
(map o apsnd) (subst_class_typ base_sort o fst o snd) operations; |
| 26518 | 349 |
val secondary_constraints = |
|
25368
f12613fda79d
proper implementation of check phase; non-qualified names for class operations
haftmann
parents:
25344
diff
changeset
|
350 |
(map o apsnd) (fn (class, (ty, _)) => subst_class_typ [class] ty) operations; |
|
25318
c8352b38d47d
synchronize_syntax: declare operations within the local scope of fixes/consts
wenzelm
parents:
25311
diff
changeset
|
351 |
fun declare_const (c, _) = |
|
c8352b38d47d
synchronize_syntax: declare operations within the local scope of fixes/consts
wenzelm
parents:
25311
diff
changeset
|
352 |
let val b = Sign.base_name c |
|
25344
00c2179db769
synchronize_syntax: improved declare_const (still inactive);
wenzelm
parents:
25326
diff
changeset
|
353 |
in Sign.intern_const thy b = c ? Variable.declare_const (b, c) end; |
| 26518 | 354 |
fun improve (c, ty) = (case AList.lookup (op =) primary_constraints c |
| 26238 | 355 |
of SOME ty' => (case try (Type.raw_match (ty', ty)) Vartab.empty |
356 |
of SOME tyenv => (case Vartab.lookup tyenv (Name.aT, 0) |
|
| 26596 | 357 |
of SOME (_, ty' as TVar (tvar as (vi, sort))) => |
| 26238 | 358 |
if TypeInfer.is_param vi |
| 26596 | 359 |
andalso Sorts.sort_le algebra (base_sort, sort) |
360 |
then SOME (ty', TFree (Name.aT, base_sort)) |
|
361 |
else NONE |
|
| 26238 | 362 |
| _ => NONE) |
363 |
| NONE => NONE) |
|
364 |
| NONE => NONE) |
|
365 |
fun subst (c, ty) = Option.map snd (AList.lookup (op =) operations c); |
|
|
25368
f12613fda79d
proper implementation of check phase; non-qualified names for class operations
haftmann
parents:
25344
diff
changeset
|
366 |
val unchecks = map (fn (c, (_, (ty, t))) => (t, Const (c, ty))) operations; |
| 25083 | 367 |
in |
368 |
ctxt |
|
| 26518 | 369 |
|> fold declare_const primary_constraints |
370 |
|> Overloading.map_improvable_syntax (K (((primary_constraints, secondary_constraints), |
|
| 26730 | 371 |
(((improve, subst), true), unchecks)), false)) |
| 26518 | 372 |
|> Overloading.set_primary_constraints |
| 25083 | 373 |
end; |
374 |
||
375 |
fun refresh_syntax class ctxt = |
|
| 25002 | 376 |
let |
377 |
val thy = ProofContext.theory_of ctxt; |
|
| 29358 | 378 |
val base_sort = base_sort thy class; |
| 26238 | 379 |
in synchronize_class_syntax [class] base_sort ctxt end; |
| 25002 | 380 |
|
| 29358 | 381 |
fun begin sups base_sort ctxt = |
| 25083 | 382 |
ctxt |
383 |
|> Variable.declare_term |
|
384 |
(Logic.mk_type (TFree (Name.aT, base_sort))) |
|
| 26238 | 385 |
|> synchronize_class_syntax sups base_sort |
386 |
|> Overloading.add_improvable_syntax; |
|
|
24901
d3cbf79769b9
added first version of user-space type system for class target
haftmann
parents:
24847
diff
changeset
|
387 |
|
| 25311 | 388 |
fun init class thy = |
389 |
thy |
|
|
29509
1ff0f3f08a7b
migrated class package to new locale implementation
haftmann
parents:
29439
diff
changeset
|
390 |
|> Locale.init class |
| 29358 | 391 |
|> begin [class] (base_sort thy class); |
|
24914
95cda5dd58d5
added proper subclass concept; improved class target
haftmann
parents:
24901
diff
changeset
|
392 |
|
| 24748 | 393 |
|
|
27690
24738db98d34
some steps towards explicit class target for canonical interpretation
haftmann
parents:
27684
diff
changeset
|
394 |
(* class target *) |
|
24738db98d34
some steps towards explicit class target for canonical interpretation
haftmann
parents:
27684
diff
changeset
|
395 |
|
|
24738db98d34
some steps towards explicit class target for canonical interpretation
haftmann
parents:
27684
diff
changeset
|
396 |
fun declare class pos ((c, mx), dict) thy = |
|
24738db98d34
some steps towards explicit class target for canonical interpretation
haftmann
parents:
27684
diff
changeset
|
397 |
let |
|
24738db98d34
some steps towards explicit class target for canonical interpretation
haftmann
parents:
27684
diff
changeset
|
398 |
val prfx = class_prefix class; |
|
24738db98d34
some steps towards explicit class target for canonical interpretation
haftmann
parents:
27684
diff
changeset
|
399 |
val thy' = thy |> Sign.add_path prfx; |
| 29439 | 400 |
val morph = morphism thy' class; |
|
28715
238f9966c80e
class morphism stemming from locale interpretation
haftmann
parents:
28674
diff
changeset
|
401 |
val inst = the_inst thy' class; |
|
238f9966c80e
class morphism stemming from locale interpretation
haftmann
parents:
28674
diff
changeset
|
402 |
val params = map (apsnd fst o snd) (these_params thy' [class]); |
|
29509
1ff0f3f08a7b
migrated class package to new locale implementation
haftmann
parents:
29439
diff
changeset
|
403 |
val amendments = map (fn class => (class, morphism thy' class)) |
|
1ff0f3f08a7b
migrated class package to new locale implementation
haftmann
parents:
29439
diff
changeset
|
404 |
(heritage thy' [class]); |
|
27690
24738db98d34
some steps towards explicit class target for canonical interpretation
haftmann
parents:
27684
diff
changeset
|
405 |
|
| 28965 | 406 |
val c' = Sign.full_bname thy' c; |
| 29439 | 407 |
val dict' = Morphism.term morph dict; |
408 |
val ty' = Term.fastype_of dict'; |
|
|
27690
24738db98d34
some steps towards explicit class target for canonical interpretation
haftmann
parents:
27684
diff
changeset
|
409 |
val ty'' = Type.strip_sorts ty'; |
| 29439 | 410 |
val def_eq = Logic.mk_equals (Const (c', ty'), dict'); |
|
29509
1ff0f3f08a7b
migrated class package to new locale implementation
haftmann
parents:
29439
diff
changeset
|
411 |
(*FIXME a mess*) |
|
1ff0f3f08a7b
migrated class package to new locale implementation
haftmann
parents:
29439
diff
changeset
|
412 |
fun amend def def' (class, morph) thy = |
|
1ff0f3f08a7b
migrated class package to new locale implementation
haftmann
parents:
29439
diff
changeset
|
413 |
Locale.amend_registration (eq_morph thy [Thm.varifyT def]) |
|
1ff0f3f08a7b
migrated class package to new locale implementation
haftmann
parents:
29439
diff
changeset
|
414 |
(class, morph) thy; |
| 28674 | 415 |
fun get_axiom thy = ((Thm.varifyT o Thm.axiom thy) c', thy); |
|
27690
24738db98d34
some steps towards explicit class target for canonical interpretation
haftmann
parents:
27684
diff
changeset
|
416 |
in |
|
24738db98d34
some steps towards explicit class target for canonical interpretation
haftmann
parents:
27684
diff
changeset
|
417 |
thy' |
| 28965 | 418 |
|> Sign.declare_const pos ((Binding.name c, ty''), mx) |> snd |
|
27690
24738db98d34
some steps towards explicit class target for canonical interpretation
haftmann
parents:
27684
diff
changeset
|
419 |
|> Thm.add_def false false (c, def_eq) |
|
24738db98d34
some steps towards explicit class target for canonical interpretation
haftmann
parents:
27684
diff
changeset
|
420 |
|>> Thm.symmetric |
|
24738db98d34
some steps towards explicit class target for canonical interpretation
haftmann
parents:
27684
diff
changeset
|
421 |
||>> get_axiom |
|
29509
1ff0f3f08a7b
migrated class package to new locale implementation
haftmann
parents:
29439
diff
changeset
|
422 |
|-> (fn (def, def') => register_operation class (c', (dict', SOME (Thm.symmetric def'))) |
|
1ff0f3f08a7b
migrated class package to new locale implementation
haftmann
parents:
29439
diff
changeset
|
423 |
#> fold (amend def def') amendments |
|
1ff0f3f08a7b
migrated class package to new locale implementation
haftmann
parents:
29439
diff
changeset
|
424 |
#> PureThy.store_thm (c ^ "_raw", def') (*FIXME name*) |
|
27690
24738db98d34
some steps towards explicit class target for canonical interpretation
haftmann
parents:
27684
diff
changeset
|
425 |
#> snd) |
|
24738db98d34
some steps towards explicit class target for canonical interpretation
haftmann
parents:
27684
diff
changeset
|
426 |
|> Sign.restore_naming thy |
|
24738db98d34
some steps towards explicit class target for canonical interpretation
haftmann
parents:
27684
diff
changeset
|
427 |
|> Sign.add_const_constraint (c', SOME ty') |
|
24738db98d34
some steps towards explicit class target for canonical interpretation
haftmann
parents:
27684
diff
changeset
|
428 |
end; |
|
24738db98d34
some steps towards explicit class target for canonical interpretation
haftmann
parents:
27684
diff
changeset
|
429 |
|
|
24738db98d34
some steps towards explicit class target for canonical interpretation
haftmann
parents:
27684
diff
changeset
|
430 |
fun abbrev class prmode pos ((c, mx), rhs) thy = |
|
24738db98d34
some steps towards explicit class target for canonical interpretation
haftmann
parents:
27684
diff
changeset
|
431 |
let |
|
24738db98d34
some steps towards explicit class target for canonical interpretation
haftmann
parents:
27684
diff
changeset
|
432 |
val prfx = class_prefix class; |
|
24738db98d34
some steps towards explicit class target for canonical interpretation
haftmann
parents:
27684
diff
changeset
|
433 |
val thy' = thy |> Sign.add_path prfx; |
|
24738db98d34
some steps towards explicit class target for canonical interpretation
haftmann
parents:
27684
diff
changeset
|
434 |
|
|
24738db98d34
some steps towards explicit class target for canonical interpretation
haftmann
parents:
27684
diff
changeset
|
435 |
val unchecks = map (fn (c, (_, (ty, t))) => (t, Const (c, ty))) |
|
24738db98d34
some steps towards explicit class target for canonical interpretation
haftmann
parents:
27684
diff
changeset
|
436 |
(these_operations thy [class]); |
| 28965 | 437 |
val c' = Sign.full_bname thy' c; |
|
27690
24738db98d34
some steps towards explicit class target for canonical interpretation
haftmann
parents:
27684
diff
changeset
|
438 |
val rhs' = Pattern.rewrite_term thy unchecks [] rhs; |
|
24738db98d34
some steps towards explicit class target for canonical interpretation
haftmann
parents:
27684
diff
changeset
|
439 |
val rhs'' = map_types Logic.varifyT rhs'; |
|
24738db98d34
some steps towards explicit class target for canonical interpretation
haftmann
parents:
27684
diff
changeset
|
440 |
val ty' = Term.fastype_of rhs'; |
|
24738db98d34
some steps towards explicit class target for canonical interpretation
haftmann
parents:
27684
diff
changeset
|
441 |
in |
|
24738db98d34
some steps towards explicit class target for canonical interpretation
haftmann
parents:
27684
diff
changeset
|
442 |
thy' |
| 28965 | 443 |
|> Sign.add_abbrev (#1 prmode) pos (Binding.name c, map_types Type.strip_sorts rhs'') |> snd |
|
27690
24738db98d34
some steps towards explicit class target for canonical interpretation
haftmann
parents:
27684
diff
changeset
|
444 |
|> Sign.add_const_constraint (c', SOME ty') |
|
24738db98d34
some steps towards explicit class target for canonical interpretation
haftmann
parents:
27684
diff
changeset
|
445 |
|> Sign.notation true prmode [(Const (c', ty'), mx)] |
|
24738db98d34
some steps towards explicit class target for canonical interpretation
haftmann
parents:
27684
diff
changeset
|
446 |
|> not (#1 prmode = PrintMode.input) ? register_operation class (c', (rhs', NONE)) |
|
24738db98d34
some steps towards explicit class target for canonical interpretation
haftmann
parents:
27684
diff
changeset
|
447 |
|> Sign.restore_naming thy |
|
24738db98d34
some steps towards explicit class target for canonical interpretation
haftmann
parents:
27684
diff
changeset
|
448 |
end; |
|
24738db98d34
some steps towards explicit class target for canonical interpretation
haftmann
parents:
27684
diff
changeset
|
449 |
|
|
24738db98d34
some steps towards explicit class target for canonical interpretation
haftmann
parents:
27684
diff
changeset
|
450 |
|
| 25462 | 451 |
(** instantiation target **) |
452 |
||
453 |
(* bookkeeping *) |
|
454 |
||
455 |
datatype instantiation = Instantiation of {
|
|
| 25864 | 456 |
arities: string list * (string * sort) list * sort, |
| 25462 | 457 |
params: ((string * string) * (string * typ)) list |
| 25603 | 458 |
(*(instantiation parameter, type constructor), (local instantiation parameter, typ)*) |
| 25462 | 459 |
} |
460 |
||
461 |
structure Instantiation = ProofDataFun |
|
462 |
( |
|
463 |
type T = instantiation |
|
| 25536 | 464 |
fun init _ = Instantiation { arities = ([], [], []), params = [] };
|
| 25462 | 465 |
); |
466 |
||
| 25485 | 467 |
fun mk_instantiation (arities, params) = |
468 |
Instantiation { arities = arities, params = params };
|
|
| 25514 | 469 |
fun get_instantiation lthy = case Instantiation.get (LocalTheory.target_of lthy) |
| 25485 | 470 |
of Instantiation data => data; |
| 25514 | 471 |
fun map_instantiation f = (LocalTheory.target o Instantiation.map) |
472 |
(fn Instantiation { arities, params } => mk_instantiation (f (arities, params)));
|
|
| 25462 | 473 |
|
| 25514 | 474 |
fun the_instantiation lthy = case get_instantiation lthy |
| 25536 | 475 |
of { arities = ([], [], []), ... } => error "No instantiation target"
|
| 25485 | 476 |
| data => data; |
| 25462 | 477 |
|
| 25485 | 478 |
val instantiation_params = #params o get_instantiation; |
| 25462 | 479 |
|
| 25514 | 480 |
fun instantiation_param lthy v = instantiation_params lthy |
| 25462 | 481 |
|> find_first (fn (_, (v', _)) => v = v') |
482 |
|> Option.map (fst o fst); |
|
483 |
||
484 |
||
485 |
(* syntax *) |
|
486 |
||
| 26238 | 487 |
fun synchronize_inst_syntax ctxt = |
| 25462 | 488 |
let |
| 26259 | 489 |
val Instantiation { arities = (_, _, sort), params = params } = Instantiation.get ctxt;
|
| 26238 | 490 |
val thy = ProofContext.theory_of ctxt; |
491 |
fun subst (c, ty) = case AxClass.inst_tyco_of thy (c, ty) |
|
492 |
of SOME tyco => (case AList.lookup (op =) params (c, tyco) |
|
493 |
of SOME (v_ty as (_, ty)) => SOME (ty, Free v_ty) |
|
494 |
| NONE => NONE) |
|
495 |
| NONE => NONE; |
|
496 |
val unchecks = |
|
497 |
map (fn ((c, _), v_ty as (_, ty)) => (Free v_ty, Const (c, ty))) params; |
|
498 |
in |
|
499 |
ctxt |
|
| 26259 | 500 |
|> Overloading.map_improvable_syntax |
| 26730 | 501 |
(fn (((primary_constraints, _), (((improve, _), _), _)), _) => |
502 |
(((primary_constraints, []), (((improve, subst), false), unchecks)), false)) |
|
| 26238 | 503 |
end; |
| 25462 | 504 |
|
505 |
||
506 |
(* target *) |
|
507 |
||
| 25485 | 508 |
val sanatize_name = (*FIXME*) |
509 |
let |
|
| 25574 | 510 |
fun is_valid s = Symbol.is_ascii_letter s orelse Symbol.is_ascii_digit s |
511 |
orelse s = "'" orelse s = "_"; |
|
| 25485 | 512 |
val is_junk = not o is_valid andf Symbol.is_regular; |
513 |
val junk = Scan.many is_junk; |
|
514 |
val scan_valids = Symbol.scanner "Malformed input" |
|
515 |
((junk |-- |
|
516 |
(Scan.optional (Scan.one Symbol.is_ascii_letter) "x" ^^ (Scan.many is_valid >> implode) |
|
517 |
--| junk)) |
|
| 25999 | 518 |
::: Scan.repeat ((Scan.many1 is_valid >> implode) --| junk)); |
| 25485 | 519 |
in |
520 |
explode #> scan_valids #> implode |
|
521 |
end; |
|
522 |
||
| 26259 | 523 |
fun type_name "*" = "prod" |
524 |
| type_name "+" = "sum" |
|
525 |
| type_name s = sanatize_name (NameSpace.base s); (*FIXME*) |
|
526 |
||
| 26518 | 527 |
fun resort_terms pp algebra consts constraints ts = |
528 |
let |
|
529 |
fun matchings (Const (c_ty as (c, _))) = (case constraints c |
|
530 |
of NONE => I |
|
531 |
| SOME sorts => fold2 (curry (Sorts.meet_sort algebra)) |
|
532 |
(Consts.typargs consts c_ty) sorts) |
|
533 |
| matchings _ = I |
|
534 |
val tvartab = (fold o fold_aterms) matchings ts Vartab.empty |
|
|
26642
454d11701fa4
Sorts.class_error: produce message only (formerly msg_class_error);
wenzelm
parents:
26628
diff
changeset
|
535 |
handle Sorts.CLASS_ERROR e => error (Sorts.class_error pp e); |
|
27089
480f19078b65
fixed wrong treatment of type variables in instantiation target
haftmann
parents:
26995
diff
changeset
|
536 |
val inst = map_type_tvar |
|
480f19078b65
fixed wrong treatment of type variables in instantiation target
haftmann
parents:
26995
diff
changeset
|
537 |
(fn (vi, sort) => TVar (vi, the_default sort (Vartab.lookup tvartab vi))); |
| 26518 | 538 |
in if Vartab.is_empty tvartab then NONE else SOME ((map o map_types) inst ts) end; |
539 |
||
| 25864 | 540 |
fun init_instantiation (tycos, vs, sort) thy = |
| 25462 | 541 |
let |
| 25536 | 542 |
val _ = if null tycos then error "At least one arity must be given" else (); |
| 26518 | 543 |
val params = these_params thy sort; |
544 |
fun get_param tyco (param, (_, (c, ty))) = if can (AxClass.param_of_inst thy) (c, tyco) |
|
| 25603 | 545 |
then NONE else SOME ((c, tyco), |
| 25864 | 546 |
(param ^ "_" ^ type_name tyco, map_atyps (K (Type (tyco, map TFree vs))) ty)); |
| 26518 | 547 |
val inst_params = map_product get_param tycos params |> map_filter I; |
548 |
val primary_constraints = map (apsnd |
|
549 |
(map_atyps (K (TVar ((Name.aT, 0), [])))) o snd o snd) params; |
|
|
26939
1035c89b4c02
moved global pretty/string_of functions from Sign to Syntax;
wenzelm
parents:
26761
diff
changeset
|
550 |
val pp = Syntax.pp_global thy; |
| 26518 | 551 |
val algebra = Sign.classes_of thy |
552 |
|> fold (fn tyco => Sorts.add_arities pp |
|
553 |
(tyco, map (fn class => (class, map snd vs)) sort)) tycos; |
|
554 |
val consts = Sign.consts_of thy; |
|
555 |
val improve_constraints = AList.lookup (op =) |
|
556 |
(map (fn (_, (class, (c, _))) => (c, [[class]])) params); |
|
557 |
fun resort_check ts ctxt = case resort_terms pp algebra consts improve_constraints ts |
|
558 |
of NONE => NONE |
|
559 |
| SOME ts' => SOME (ts', ctxt); |
|
560 |
fun improve (c, ty) = case AxClass.inst_tyco_of thy (c, ty) |
|
|
26329
3e58e4c67a2a
instantiation less liberal with dangling constraints
haftmann
parents:
26259
diff
changeset
|
561 |
of SOME tyco => (case AList.lookup (op =) inst_params (c, tyco) |
| 26518 | 562 |
of SOME (_, ty') => if Type.raw_instance (ty', ty) then SOME (ty, ty') else NONE |
| 26259 | 563 |
| NONE => NONE) |
564 |
| NONE => NONE; |
|
| 25485 | 565 |
in |
566 |
thy |
|
567 |
|> ProofContext.init |
|
|
26329
3e58e4c67a2a
instantiation less liberal with dangling constraints
haftmann
parents:
26259
diff
changeset
|
568 |
|> Instantiation.put (mk_instantiation ((tycos, vs, sort), inst_params)) |
| 27281 | 569 |
|> fold (Variable.declare_typ o TFree) vs |
|
26329
3e58e4c67a2a
instantiation less liberal with dangling constraints
haftmann
parents:
26259
diff
changeset
|
570 |
|> fold (Variable.declare_names o Free o snd) inst_params |
| 26259 | 571 |
|> (Overloading.map_improvable_syntax o apfst) |
|
26329
3e58e4c67a2a
instantiation less liberal with dangling constraints
haftmann
parents:
26259
diff
changeset
|
572 |
(fn ((_, _), ((_, subst), unchecks)) => |
| 26730 | 573 |
((primary_constraints, []), (((improve, K NONE), false), []))) |
| 26259 | 574 |
|> Overloading.add_improvable_syntax |
| 26518 | 575 |
|> Context.proof_map (Syntax.add_term_check 0 "resorting" resort_check) |
| 26238 | 576 |
|> synchronize_inst_syntax |
| 25485 | 577 |
end; |
578 |
||
| 26238 | 579 |
fun confirm_declaration c = (map_instantiation o apsnd) |
580 |
(filter_out (fn (_, (c', _)) => c' = c)) |
|
581 |
#> LocalTheory.target synchronize_inst_syntax |
|
582 |
||
| 25485 | 583 |
fun gen_instantiation_instance do_proof after_qed lthy = |
584 |
let |
|
| 25864 | 585 |
val (tycos, vs, sort) = (#arities o the_instantiation) lthy; |
586 |
val arities_proof = maps (fn tyco => Logic.mk_arities (tyco, map snd vs, sort)) tycos; |
|
| 25462 | 587 |
fun after_qed' results = |
588 |
LocalTheory.theory (fold (AxClass.add_arity o Thm.varifyT) results) |
|
589 |
#> after_qed; |
|
590 |
in |
|
591 |
lthy |
|
592 |
|> do_proof after_qed' arities_proof |
|
593 |
end; |
|
594 |
||
| 25485 | 595 |
val instantiation_instance = gen_instantiation_instance (fn after_qed => fn ts => |
| 25462 | 596 |
Proof.theorem_i NONE (after_qed o map the_single) (map (fn t => [(t, [])]) ts)); |
597 |
||
| 25485 | 598 |
fun prove_instantiation_instance tac = gen_instantiation_instance (fn after_qed => |
| 25502 | 599 |
fn ts => fn lthy => after_qed (map (fn t => Goal.prove lthy [] [] t |
600 |
(fn {context, ...} => tac context)) ts) lthy) I;
|
|
| 25462 | 601 |
|
| 28666 | 602 |
fun prove_instantiation_exit tac = prove_instantiation_instance tac |
603 |
#> LocalTheory.exit_global; |
|
604 |
||
605 |
fun prove_instantiation_exit_result f tac x lthy = |
|
606 |
let |
|
| 29439 | 607 |
val morph = ProofContext.export_morphism lthy |
| 28666 | 608 |
(ProofContext.init (ProofContext.theory_of lthy)); |
| 29439 | 609 |
val y = f morph x; |
| 28666 | 610 |
in |
611 |
lthy |
|
612 |
|> prove_instantiation_exit (fn ctxt => tac ctxt y) |
|
613 |
|> pair y |
|
614 |
end; |
|
615 |
||
| 25462 | 616 |
fun conclude_instantiation lthy = |
617 |
let |
|
| 25485 | 618 |
val { arities, params } = the_instantiation lthy;
|
| 25864 | 619 |
val (tycos, vs, sort) = arities; |
| 25462 | 620 |
val thy = ProofContext.theory_of lthy; |
|
25597
34860182b250
moved instance parameter management from class.ML to axclass.ML
haftmann
parents:
25574
diff
changeset
|
621 |
val _ = map (fn tyco => if Sign.of_sort thy |
| 25864 | 622 |
(Type (tyco, map TFree vs), sort) |
| 25462 | 623 |
then () else error ("Missing instance proof for type " ^ quote (Sign.extern_type thy tyco)))
|
|
25597
34860182b250
moved instance parameter management from class.ML to axclass.ML
haftmann
parents:
25574
diff
changeset
|
624 |
tycos; |
|
34860182b250
moved instance parameter management from class.ML to axclass.ML
haftmann
parents:
25574
diff
changeset
|
625 |
in lthy end; |
| 25462 | 626 |
|
| 25603 | 627 |
fun pretty_instantiation lthy = |
628 |
let |
|
629 |
val { arities, params } = the_instantiation lthy;
|
|
| 25864 | 630 |
val (tycos, vs, sort) = arities; |
| 25603 | 631 |
val thy = ProofContext.theory_of lthy; |
| 25864 | 632 |
fun pr_arity tyco = Syntax.pretty_arity lthy (tyco, map snd vs, sort); |
| 25603 | 633 |
fun pr_param ((c, _), (v, ty)) = |
| 25864 | 634 |
(Pretty.block o Pretty.breaks) [Pretty.str v, Pretty.str "==", |
|
26939
1035c89b4c02
moved global pretty/string_of functions from Sign to Syntax;
wenzelm
parents:
26761
diff
changeset
|
635 |
(Pretty.str o Sign.extern_const thy) c, Pretty.str "::", Syntax.pretty_typ_global thy ty]; |
| 25603 | 636 |
in |
637 |
(Pretty.block o Pretty.fbreaks) |
|
638 |
(Pretty.str "instantiation" :: map pr_arity tycos @ map pr_param params) |
|
639 |
end; |
|
640 |
||
| 24218 | 641 |
end; |
| 25683 | 642 |