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