haftmann@24218
|
1 |
(* Title: Pure/Isar/class.ML
|
haftmann@24218
|
2 |
ID: $Id$
|
haftmann@24218
|
3 |
Author: Florian Haftmann, TU Muenchen
|
haftmann@24218
|
4 |
|
haftmann@24218
|
5 |
Type classes derived from primitive axclasses and locales.
|
haftmann@24218
|
6 |
*)
|
haftmann@24218
|
7 |
|
haftmann@24218
|
8 |
signature CLASS =
|
haftmann@24218
|
9 |
sig
|
haftmann@25462
|
10 |
(*classes*)
|
haftmann@25002
|
11 |
val class: bstring -> class list -> Element.context_i Locale.element list
|
haftmann@24218
|
12 |
-> string list -> theory -> string * Proof.context
|
haftmann@25002
|
13 |
val class_cmd: bstring -> xstring list -> Element.context Locale.element list
|
haftmann@24589
|
14 |
-> xstring list -> theory -> string * Proof.context
|
haftmann@25485
|
15 |
|
haftmann@25311
|
16 |
val init: class -> theory -> Proof.context
|
haftmann@25603
|
17 |
val declare: string -> Markup.property list
|
wenzelm@25104
|
18 |
-> (string * mixfix) * term -> theory -> theory
|
haftmann@25603
|
19 |
val abbrev: string -> Syntax.mode -> Markup.property list
|
wenzelm@25104
|
20 |
-> (string * mixfix) * term -> theory -> theory
|
haftmann@25083
|
21 |
val refresh_syntax: class -> Proof.context -> Proof.context
|
haftmann@25485
|
22 |
|
haftmann@24589
|
23 |
val intro_classes_tac: thm list -> tactic
|
haftmann@24589
|
24 |
val default_intro_classes_tac: thm list -> tactic
|
haftmann@25618
|
25 |
val prove_subclass: class * class -> thm -> theory -> theory
|
haftmann@25485
|
26 |
|
haftmann@25485
|
27 |
val class_prefix: string -> string
|
haftmann@25485
|
28 |
val is_class: theory -> class -> bool
|
haftmann@25485
|
29 |
val these_params: theory -> sort -> (string * (string * typ)) list
|
haftmann@24589
|
30 |
val print_classes: theory -> unit
|
haftmann@24423
|
31 |
|
haftmann@25462
|
32 |
(*instances*)
|
haftmann@25864
|
33 |
val init_instantiation: string list * (string * sort) list * sort -> theory -> local_theory
|
haftmann@25485
|
34 |
val instantiation_instance: (local_theory -> local_theory) -> local_theory -> Proof.state
|
haftmann@25485
|
35 |
val prove_instantiation_instance: (Proof.context -> tactic) -> local_theory -> local_theory
|
haftmann@25485
|
36 |
val conclude_instantiation: local_theory -> local_theory
|
haftmann@25603
|
37 |
val instantiation_param: local_theory -> string -> string option
|
haftmann@25485
|
38 |
val confirm_declaration: string -> local_theory -> local_theory
|
haftmann@25603
|
39 |
val pretty_instantiation: local_theory -> Pretty.T
|
haftmann@25485
|
40 |
|
haftmann@25462
|
41 |
(*old axclass layer*)
|
haftmann@25462
|
42 |
val axclass_cmd: bstring * xstring list
|
haftmann@25462
|
43 |
-> ((bstring * Attrib.src list) * string list) list
|
haftmann@25462
|
44 |
-> theory -> class * theory
|
haftmann@25462
|
45 |
val classrel_cmd: xstring * xstring -> theory -> Proof.state
|
haftmann@25462
|
46 |
|
haftmann@25462
|
47 |
(*old instance layer*)
|
haftmann@25536
|
48 |
val instance_arity: (theory -> theory) -> arity -> theory -> Proof.state
|
haftmann@25536
|
49 |
val instance_arity_cmd: bstring * xstring list * xstring -> theory -> Proof.state
|
haftmann@24218
|
50 |
end;
|
haftmann@24218
|
51 |
|
haftmann@24218
|
52 |
structure Class : CLASS =
|
haftmann@24218
|
53 |
struct
|
haftmann@24218
|
54 |
|
haftmann@24218
|
55 |
(** auxiliary **)
|
haftmann@24218
|
56 |
|
haftmann@25062
|
57 |
val classN = "class";
|
haftmann@25062
|
58 |
val introN = "intro";
|
haftmann@25062
|
59 |
|
haftmann@25002
|
60 |
fun prove_interpretation tac prfx_atts expr inst =
|
haftmann@25002
|
61 |
Locale.interpretation_i I prfx_atts expr inst
|
haftmann@24589
|
62 |
#> Proof.global_terminal_proof
|
haftmann@24589
|
63 |
(Method.Basic (K (Method.SIMPLE_METHOD tac), Position.none), NONE)
|
haftmann@24589
|
64 |
#> ProofContext.theory_of;
|
haftmann@24589
|
65 |
|
haftmann@25195
|
66 |
fun prove_interpretation_in tac after_qed (name, expr) =
|
haftmann@25195
|
67 |
Locale.interpretation_in_locale
|
haftmann@25195
|
68 |
(ProofContext.theory after_qed) (name, expr)
|
haftmann@25195
|
69 |
#> Proof.global_terminal_proof
|
haftmann@25195
|
70 |
(Method.Basic (K (Method.SIMPLE_METHOD tac), Position.none), NONE)
|
haftmann@25195
|
71 |
#> ProofContext.theory_of;
|
haftmann@25195
|
72 |
|
haftmann@25038
|
73 |
fun get_remove_global_constraint c thy =
|
haftmann@25038
|
74 |
let
|
haftmann@25038
|
75 |
val ty = Sign.the_const_constraint thy c;
|
haftmann@25038
|
76 |
in
|
haftmann@25038
|
77 |
thy
|
haftmann@25038
|
78 |
|> Sign.add_const_constraint (c, NONE)
|
haftmann@25038
|
79 |
|> pair (c, Logic.unvarifyT ty)
|
haftmann@25038
|
80 |
end;
|
haftmann@25038
|
81 |
|
haftmann@24589
|
82 |
|
haftmann@25485
|
83 |
(** primitive axclass and instance commands **)
|
haftmann@24589
|
84 |
|
haftmann@24218
|
85 |
fun axclass_cmd (class, raw_superclasses) raw_specs thy =
|
haftmann@24218
|
86 |
let
|
haftmann@24218
|
87 |
val ctxt = ProofContext.init thy;
|
haftmann@24218
|
88 |
val superclasses = map (Sign.read_class thy) raw_superclasses;
|
haftmann@24589
|
89 |
val name_atts = map ((apsnd o map) (Attrib.attribute thy) o fst)
|
haftmann@24589
|
90 |
raw_specs;
|
haftmann@24589
|
91 |
val axiomss = ProofContext.read_propp (ctxt, map (map (rpair []) o snd)
|
haftmann@24589
|
92 |
raw_specs)
|
haftmann@24218
|
93 |
|> snd
|
haftmann@24218
|
94 |
|> (map o map) fst;
|
haftmann@24589
|
95 |
in
|
haftmann@24589
|
96 |
AxClass.define_class (class, superclasses) []
|
haftmann@24589
|
97 |
(name_atts ~~ axiomss) thy
|
haftmann@24589
|
98 |
end;
|
haftmann@24218
|
99 |
|
haftmann@24218
|
100 |
local
|
haftmann@24218
|
101 |
|
haftmann@24218
|
102 |
fun gen_instance mk_prop add_thm after_qed insts thy =
|
haftmann@24218
|
103 |
let
|
haftmann@24218
|
104 |
fun after_qed' results =
|
haftmann@24218
|
105 |
ProofContext.theory ((fold o fold) add_thm results #> after_qed);
|
haftmann@24218
|
106 |
in
|
haftmann@24218
|
107 |
thy
|
haftmann@24218
|
108 |
|> ProofContext.init
|
haftmann@24589
|
109 |
|> Proof.theorem_i NONE after_qed' ((map (fn t => [(t, [])])
|
haftmann@25536
|
110 |
o mk_prop thy) insts)
|
haftmann@24218
|
111 |
end;
|
haftmann@24218
|
112 |
|
haftmann@24218
|
113 |
in
|
haftmann@24218
|
114 |
|
haftmann@24589
|
115 |
val instance_arity =
|
haftmann@24218
|
116 |
gen_instance (Logic.mk_arities oo Sign.cert_arity) AxClass.add_arity;
|
haftmann@25502
|
117 |
val instance_arity_cmd =
|
haftmann@25502
|
118 |
gen_instance (Logic.mk_arities oo Sign.read_arity) AxClass.add_arity I;
|
haftmann@24589
|
119 |
val classrel =
|
haftmann@25536
|
120 |
gen_instance (single oo (Logic.mk_classrel oo AxClass.cert_classrel)) AxClass.add_classrel I;
|
haftmann@24589
|
121 |
val classrel_cmd =
|
haftmann@25536
|
122 |
gen_instance (single oo (Logic.mk_classrel oo AxClass.read_classrel)) AxClass.add_classrel I;
|
haftmann@24218
|
123 |
|
haftmann@24218
|
124 |
end; (*local*)
|
haftmann@24218
|
125 |
|
haftmann@24218
|
126 |
|
haftmann@24589
|
127 |
(** class data **)
|
haftmann@24218
|
128 |
|
haftmann@24218
|
129 |
datatype class_data = ClassData of {
|
haftmann@24218
|
130 |
consts: (string * string) list
|
haftmann@24836
|
131 |
(*locale parameter ~> constant name*),
|
haftmann@25062
|
132 |
base_sort: sort,
|
haftmann@25083
|
133 |
inst: term option list
|
haftmann@25083
|
134 |
(*canonical interpretation*),
|
haftmann@25711
|
135 |
morphism: theory -> thm list -> morphism,
|
haftmann@25062
|
136 |
(*partial morphism of canonical interpretation*)
|
haftmann@25618
|
137 |
assm_intro: thm option,
|
haftmann@25618
|
138 |
of_class: thm,
|
haftmann@25618
|
139 |
axiom: thm option,
|
haftmann@24657
|
140 |
defs: thm list,
|
haftmann@25368
|
141 |
operations: (string * (class * (typ * term))) list
|
haftmann@24657
|
142 |
};
|
haftmann@24218
|
143 |
|
haftmann@24657
|
144 |
fun rep_class_data (ClassData d) = d;
|
haftmann@25618
|
145 |
fun mk_class_data ((consts, base_sort, inst, morphism, assm_intro, of_class, axiom),
|
haftmann@25368
|
146 |
(defs, operations)) =
|
haftmann@25062
|
147 |
ClassData { consts = consts, base_sort = base_sort, inst = inst,
|
haftmann@25618
|
148 |
morphism = morphism, assm_intro = assm_intro, of_class = of_class, axiom = axiom,
|
haftmann@25618
|
149 |
defs = defs, operations = operations };
|
haftmann@25618
|
150 |
fun map_class_data f (ClassData { consts, base_sort, inst, morphism,
|
haftmann@25618
|
151 |
assm_intro, of_class, axiom, defs, operations }) =
|
haftmann@25618
|
152 |
mk_class_data (f ((consts, base_sort, inst, morphism, assm_intro, of_class, axiom),
|
haftmann@25368
|
153 |
(defs, operations)));
|
haftmann@25038
|
154 |
fun merge_class_data _ (ClassData { consts = consts,
|
haftmann@25618
|
155 |
base_sort = base_sort, inst = inst, morphism = morphism, assm_intro = assm_intro,
|
haftmann@25618
|
156 |
of_class = of_class, axiom = axiom, defs = defs1, operations = operations1 },
|
haftmann@25618
|
157 |
ClassData { consts = _, base_sort = _, inst = _, morphism = _, assm_intro = _,
|
haftmann@25618
|
158 |
of_class = _, axiom = _, defs = defs2, operations = operations2 }) =
|
haftmann@25618
|
159 |
mk_class_data ((consts, base_sort, inst, morphism, assm_intro, of_class, axiom),
|
haftmann@24914
|
160 |
(Thm.merge_thms (defs1, defs2),
|
haftmann@25368
|
161 |
AList.merge (op =) (K true) (operations1, operations2)));
|
haftmann@24218
|
162 |
|
haftmann@24218
|
163 |
structure ClassData = TheoryDataFun
|
haftmann@24218
|
164 |
(
|
haftmann@25038
|
165 |
type T = class_data Graph.T
|
haftmann@25038
|
166 |
val empty = Graph.empty;
|
haftmann@24218
|
167 |
val copy = I;
|
haftmann@24218
|
168 |
val extend = I;
|
haftmann@25038
|
169 |
fun merge _ = Graph.join merge_class_data;
|
haftmann@24218
|
170 |
);
|
haftmann@24218
|
171 |
|
haftmann@24218
|
172 |
|
haftmann@24218
|
173 |
(* queries *)
|
haftmann@24218
|
174 |
|
haftmann@25038
|
175 |
val lookup_class_data = Option.map rep_class_data oo try o Graph.get_node o ClassData.get;
|
haftmann@24218
|
176 |
|
haftmann@24589
|
177 |
fun the_class_data thy class = case lookup_class_data thy class
|
wenzelm@25020
|
178 |
of NONE => error ("Undeclared class " ^ quote class)
|
haftmann@24589
|
179 |
| SOME data => data;
|
haftmann@24218
|
180 |
|
haftmann@25038
|
181 |
val is_class = is_some oo lookup_class_data;
|
haftmann@25038
|
182 |
|
haftmann@25038
|
183 |
val ancestry = Graph.all_succs o ClassData.get;
|
haftmann@24218
|
184 |
|
haftmann@25002
|
185 |
fun these_params thy =
|
haftmann@24218
|
186 |
let
|
haftmann@24218
|
187 |
fun params class =
|
haftmann@24218
|
188 |
let
|
wenzelm@24930
|
189 |
val const_typs = (#params o AxClass.get_info thy) class;
|
haftmann@24657
|
190 |
val const_names = (#consts o the_class_data thy) class;
|
haftmann@24218
|
191 |
in
|
haftmann@24218
|
192 |
(map o apsnd) (fn c => (c, (the o AList.lookup (op =) const_typs) c)) const_names
|
haftmann@24218
|
193 |
end;
|
haftmann@24218
|
194 |
in maps params o ancestry thy end;
|
haftmann@24218
|
195 |
|
haftmann@24657
|
196 |
fun these_defs thy = maps (these o Option.map #defs o lookup_class_data thy) o ancestry thy;
|
haftmann@24218
|
197 |
|
haftmann@25711
|
198 |
fun partial_morphism thy class = #morphism (the_class_data thy class) thy [];
|
haftmann@25711
|
199 |
fun morphism thy class = #morphism (the_class_data thy class) thy (these_defs thy [class]);
|
haftmann@25062
|
200 |
|
haftmann@25618
|
201 |
fun these_assm_intros thy =
|
haftmann@25618
|
202 |
Graph.fold (fn (_, (data, _)) => fold (insert Thm.eq_thm)
|
haftmann@25618
|
203 |
((the_list o #assm_intro o rep_class_data) data)) (ClassData.get thy) [];
|
haftmann@24218
|
204 |
|
haftmann@24836
|
205 |
fun these_operations thy =
|
haftmann@24836
|
206 |
maps (#operations o the_class_data thy) o ancestry thy;
|
haftmann@24657
|
207 |
|
haftmann@24218
|
208 |
fun print_classes thy =
|
haftmann@24218
|
209 |
let
|
wenzelm@24920
|
210 |
val ctxt = ProofContext.init thy;
|
haftmann@24218
|
211 |
val algebra = Sign.classes_of thy;
|
haftmann@24218
|
212 |
val arities =
|
haftmann@24218
|
213 |
Symtab.empty
|
haftmann@24218
|
214 |
|> Symtab.fold (fn (tyco, arities) => fold (fn (class, _) =>
|
haftmann@24218
|
215 |
Symtab.map_default (class, []) (insert (op =) tyco)) arities)
|
haftmann@24218
|
216 |
((#arities o Sorts.rep_algebra) algebra);
|
haftmann@24218
|
217 |
val the_arities = these o Symtab.lookup arities;
|
haftmann@24218
|
218 |
fun mk_arity class tyco =
|
haftmann@24218
|
219 |
let
|
haftmann@24218
|
220 |
val Ss = Sorts.mg_domain algebra tyco [class];
|
wenzelm@24920
|
221 |
in Syntax.pretty_arity ctxt (tyco, Ss, [class]) end;
|
haftmann@24218
|
222 |
fun mk_param (c, ty) = Pretty.str (Sign.extern_const thy c ^ " :: "
|
wenzelm@24920
|
223 |
^ setmp show_sorts false (Syntax.string_of_typ ctxt o Type.strip_sorts) ty);
|
haftmann@24218
|
224 |
fun mk_entry class = (Pretty.block o Pretty.fbreaks o map_filter I) [
|
haftmann@25062
|
225 |
(SOME o Pretty.str) ("class " ^ Sign.extern_class thy class ^ ":"),
|
haftmann@24218
|
226 |
(SOME o Pretty.block) [Pretty.str "supersort: ",
|
wenzelm@24920
|
227 |
(Syntax.pretty_sort ctxt o Sign.minimize_sort thy o Sign.super_classes thy) class],
|
haftmann@25062
|
228 |
if is_class thy class then (SOME o Pretty.str)
|
haftmann@25062
|
229 |
("locale: " ^ Locale.extern thy class) else NONE,
|
haftmann@25062
|
230 |
((fn [] => NONE | ps => (SOME o Pretty.block o Pretty.fbreaks)
|
haftmann@25062
|
231 |
(Pretty.str "parameters:" :: ps)) o map mk_param
|
wenzelm@24930
|
232 |
o these o Option.map #params o try (AxClass.get_info thy)) class,
|
haftmann@24218
|
233 |
(SOME o Pretty.block o Pretty.breaks) [
|
haftmann@24218
|
234 |
Pretty.str "instances:",
|
haftmann@24218
|
235 |
Pretty.list "" "" (map (mk_arity class) (the_arities class))
|
haftmann@24218
|
236 |
]
|
haftmann@24218
|
237 |
]
|
haftmann@24218
|
238 |
in
|
haftmann@24589
|
239 |
(Pretty.writeln o Pretty.chunks o separate (Pretty.str "")
|
haftmann@24589
|
240 |
o map mk_entry o Sorts.all_classes) algebra
|
haftmann@24218
|
241 |
end;
|
haftmann@24218
|
242 |
|
haftmann@24218
|
243 |
|
haftmann@24218
|
244 |
(* updaters *)
|
haftmann@24218
|
245 |
|
haftmann@25618
|
246 |
fun add_class_data ((class, superclasses),
|
haftmann@25711
|
247 |
(params, base_sort, inst, phi, axiom, assm_intro, of_class)) thy =
|
haftmann@25002
|
248 |
let
|
haftmann@25368
|
249 |
val operations = map (fn (v_ty as (_, ty), (c, _)) =>
|
haftmann@25683
|
250 |
(c, (class, (ty, Free v_ty)))) params;
|
haftmann@25038
|
251 |
val add_class = Graph.new_node (class,
|
haftmann@25683
|
252 |
mk_class_data (((map o pairself) fst params, base_sort,
|
haftmann@25618
|
253 |
map (SOME o Const) inst, phi, assm_intro, of_class, axiom), ([], operations)))
|
haftmann@25002
|
254 |
#> fold (curry Graph.add_edge class) superclasses;
|
haftmann@25618
|
255 |
in ClassData.map add_class thy end;
|
haftmann@24218
|
256 |
|
haftmann@25368
|
257 |
fun register_operation class (c, (t, some_def)) thy =
|
haftmann@25062
|
258 |
let
|
haftmann@25368
|
259 |
val base_sort = (#base_sort o the_class_data thy) class;
|
haftmann@25239
|
260 |
val prep_typ = map_atyps
|
haftmann@25368
|
261 |
(fn TVar (vi as (v, _), sort) => if Name.aT = v
|
haftmann@25368
|
262 |
then TFree (v, base_sort) else TVar (vi, sort));
|
haftmann@25368
|
263 |
val t' = map_types prep_typ t;
|
haftmann@25368
|
264 |
val ty' = Term.fastype_of t';
|
haftmann@25062
|
265 |
in
|
haftmann@25062
|
266 |
thy
|
haftmann@25062
|
267 |
|> (ClassData.map o Graph.map_node class o map_class_data o apsnd)
|
haftmann@25368
|
268 |
(fn (defs, operations) =>
|
haftmann@25096
|
269 |
(fold cons (the_list some_def) defs,
|
haftmann@25368
|
270 |
(c, (class, (ty', t'))) :: operations))
|
haftmann@25062
|
271 |
end;
|
haftmann@24218
|
272 |
|
haftmann@24589
|
273 |
|
haftmann@24589
|
274 |
(** rule calculation, tactics and methods **)
|
haftmann@24589
|
275 |
|
wenzelm@25024
|
276 |
val class_prefix = Logic.const_of_class o Sign.base_name;
|
wenzelm@25024
|
277 |
|
haftmann@25683
|
278 |
fun calculate thy sups base_sort assm_axiom param_map class =
|
haftmann@25062
|
279 |
let
|
haftmann@25711
|
280 |
(*static parts of morphism*)
|
haftmann@25683
|
281 |
val subst_typ = map_atyps (fn TFree (v, sort) =>
|
haftmann@25683
|
282 |
if v = Name.aT then TVar ((v, 0), [class]) else TVar ((v, 0), sort)
|
haftmann@25683
|
283 |
| ty => ty);
|
haftmann@25683
|
284 |
fun subst_aterm (t as Free (v, ty)) = (case AList.lookup (op =) param_map v
|
haftmann@25062
|
285 |
of SOME (c, _) => Const (c, ty)
|
haftmann@25062
|
286 |
| NONE => t)
|
haftmann@25062
|
287 |
| subst_aterm t = t;
|
haftmann@25711
|
288 |
fun instantiate thy sort = Thm.instantiate ([pairself (Thm.ctyp_of thy o TVar o pair (Name.aT, 0))
|
haftmann@25711
|
289 |
(base_sort, sort)], map (fn (v, (c, ty)) => pairself (Thm.cterm_of thy)
|
haftmann@25711
|
290 |
(Var ((v, 0), map_atyps (fn _ => TVar ((Name.aT, 0), sort)) ty),
|
haftmann@25711
|
291 |
Const (c, map_atyps (fn _ => TVar ((Name.aT, 0), sort)) ty))) param_map);
|
haftmann@25711
|
292 |
val instantiate_base_sort = instantiate thy base_sort;
|
haftmann@25711
|
293 |
val instantiate_class = instantiate thy [class];
|
haftmann@25683
|
294 |
val (proto_assm_intro, locale_intro) = Locale.intros thy class
|
haftmann@25683
|
295 |
|> pairself (try the_single);
|
haftmann@25683
|
296 |
val axiom_premises = map_filter (#axiom o the_class_data thy) sups
|
haftmann@25683
|
297 |
@ the_list assm_axiom;
|
haftmann@25711
|
298 |
val axiom = locale_intro
|
haftmann@25711
|
299 |
|> Option.map (Drule.standard o (fn thm => thm OF axiom_premises) o instantiate_class)
|
haftmann@25711
|
300 |
|> (fn x as SOME _ => x | NONE => assm_axiom);
|
haftmann@25711
|
301 |
val lift_axiom = case axiom
|
haftmann@25711
|
302 |
of SOME axiom => (fn thm => Thm.implies_elim thm axiom)
|
haftmann@25683
|
303 |
| NONE => I;
|
haftmann@25062
|
304 |
|
haftmann@25711
|
305 |
(*dynamic parts of morphism*)
|
haftmann@25711
|
306 |
fun rew_term thy defs = Pattern.rewrite_term thy
|
haftmann@25711
|
307 |
(map (Logic.dest_equals o Thm.prop_of) defs) [];
|
haftmann@25711
|
308 |
fun subst_term thy defs = map_aterms subst_aterm #> rew_term thy defs
|
haftmann@25711
|
309 |
#> map_types subst_typ;
|
haftmann@25711
|
310 |
fun subst_thm defs = Drule.standard' #> instantiate_class #> lift_axiom
|
haftmann@25711
|
311 |
#> MetaSimplifier.rewrite_rule defs;
|
haftmann@25711
|
312 |
fun morphism thy defs =
|
haftmann@25711
|
313 |
Morphism.typ_morphism subst_typ
|
haftmann@25711
|
314 |
$> Morphism.term_morphism (subst_term thy defs)
|
haftmann@25711
|
315 |
$> Morphism.thm_morphism (subst_thm defs);
|
haftmann@25711
|
316 |
|
haftmann@25711
|
317 |
(*class rules*)
|
haftmann@25711
|
318 |
val defs = these_defs thy sups;
|
haftmann@25618
|
319 |
val assm_intro = proto_assm_intro
|
haftmann@25711
|
320 |
|> Option.map instantiate_base_sort
|
haftmann@25711
|
321 |
|> Option.map (MetaSimplifier.rewrite_rule defs)
|
haftmann@25668
|
322 |
|> Option.map Goal.close_result;
|
haftmann@25711
|
323 |
val fixate = Thm.instantiate
|
haftmann@25711
|
324 |
(map (pairself (Thm.ctyp_of thy)) [(TVar ((Name.aT, 0), []), TFree (Name.aT, base_sort)),
|
haftmann@25711
|
325 |
(TVar ((Name.aT, 0), base_sort), TFree (Name.aT, base_sort))], [])
|
haftmann@25711
|
326 |
val class_intro = (fixate o #intro o AxClass.get_info thy) class;
|
haftmann@25618
|
327 |
val of_class_sups = if null sups
|
haftmann@25711
|
328 |
then map (fixate o Thm.class_triv thy) base_sort
|
haftmann@25711
|
329 |
else map (fixate o #of_class o the_class_data thy) sups;
|
haftmann@25683
|
330 |
val locale_dests = map Drule.standard' (Locale.dests thy class);
|
haftmann@25711
|
331 |
val num_trivs = case length locale_dests
|
haftmann@25711
|
332 |
of 0 => if is_none axiom then 0 else 1
|
haftmann@25711
|
333 |
| n => n;
|
haftmann@25711
|
334 |
val pred_trivs = if num_trivs = 0 then []
|
haftmann@25711
|
335 |
else the axiom
|
haftmann@25711
|
336 |
|> Thm.prop_of
|
haftmann@25711
|
337 |
|> (map_types o map_atyps o K) (TFree (Name.aT, base_sort))
|
haftmann@25711
|
338 |
|> (Thm.assume o Thm.cterm_of thy)
|
haftmann@25711
|
339 |
|> replicate num_trivs;
|
haftmann@25668
|
340 |
val of_class = (class_intro OF of_class_sups OF locale_dests OF pred_trivs)
|
haftmann@25711
|
341 |
|> Drule.standard'
|
haftmann@25668
|
342 |
|> Goal.close_result;
|
haftmann@25711
|
343 |
in (morphism, axiom, assm_intro, of_class) end;
|
haftmann@24589
|
344 |
|
haftmann@24589
|
345 |
fun class_interpretation class facts defs thy =
|
haftmann@24589
|
346 |
let
|
haftmann@25038
|
347 |
val params = these_params thy [class];
|
haftmann@25083
|
348 |
val inst = (#inst o the_class_data thy) class;
|
wenzelm@25020
|
349 |
val tac = ALLGOALS (ProofContext.fact_tac facts);
|
haftmann@25038
|
350 |
val prfx = class_prefix class;
|
haftmann@24589
|
351 |
in
|
haftmann@25038
|
352 |
thy
|
haftmann@25038
|
353 |
|> fold_map (get_remove_global_constraint o fst o snd) params
|
ballarin@25094
|
354 |
||> prove_interpretation tac ((false, prfx), []) (Locale.Locale class)
|
ballarin@25094
|
355 |
(inst, map (fn def => (("", []), def)) defs)
|
haftmann@25038
|
356 |
|-> (fn cs => fold (Sign.add_const_constraint o apsnd SOME) cs)
|
haftmann@24589
|
357 |
end;
|
haftmann@24218
|
358 |
|
haftmann@25618
|
359 |
fun prove_subclass (sub, sup) thm thy =
|
haftmann@25618
|
360 |
let
|
haftmann@25711
|
361 |
val of_class = (#of_class o the_class_data thy) sup;
|
haftmann@25683
|
362 |
val intro = Drule.standard' (of_class OF [Drule.standard' thm]);
|
haftmann@25618
|
363 |
val classrel = intro OF (the_list o #axiom o the_class_data thy) sub;
|
haftmann@25618
|
364 |
in
|
haftmann@25618
|
365 |
thy
|
haftmann@25618
|
366 |
|> AxClass.add_classrel classrel
|
haftmann@25618
|
367 |
|> prove_interpretation_in (ALLGOALS (ProofContext.fact_tac [thm]))
|
haftmann@25618
|
368 |
I (sub, Locale.Locale sup)
|
haftmann@25618
|
369 |
|> ClassData.map (Graph.add_edge (sub, sup))
|
haftmann@25618
|
370 |
end;
|
haftmann@25618
|
371 |
|
haftmann@24218
|
372 |
fun intro_classes_tac facts st =
|
haftmann@24218
|
373 |
let
|
haftmann@24218
|
374 |
val thy = Thm.theory_of_thm st;
|
haftmann@24218
|
375 |
val classes = Sign.all_classes thy;
|
haftmann@24218
|
376 |
val class_trivs = map (Thm.class_triv thy) classes;
|
haftmann@25618
|
377 |
val class_intros = map_filter (try (#intro o AxClass.get_info thy)) classes;
|
haftmann@25618
|
378 |
val assm_intros = these_assm_intros thy;
|
haftmann@24218
|
379 |
in
|
haftmann@25618
|
380 |
Method.intros_tac (class_trivs @ class_intros @ assm_intros) facts st
|
haftmann@24218
|
381 |
end;
|
haftmann@24218
|
382 |
|
haftmann@24218
|
383 |
fun default_intro_classes_tac [] = intro_classes_tac []
|
wenzelm@24930
|
384 |
| default_intro_classes_tac _ = no_tac;
|
haftmann@24218
|
385 |
|
haftmann@24218
|
386 |
fun default_tac rules ctxt facts =
|
haftmann@24218
|
387 |
HEADGOAL (Method.some_rule_tac rules ctxt facts) ORELSE
|
haftmann@24218
|
388 |
default_intro_classes_tac facts;
|
haftmann@24218
|
389 |
|
haftmann@24218
|
390 |
val _ = Context.add_setup (Method.add_methods
|
haftmann@24218
|
391 |
[("intro_classes", Method.no_args (Method.METHOD intro_classes_tac),
|
haftmann@24218
|
392 |
"back-chain introduction rules of classes"),
|
haftmann@24218
|
393 |
("default", Method.thms_ctxt_args (Method.METHOD oo default_tac),
|
haftmann@24218
|
394 |
"apply some intro/elim rule")]);
|
haftmann@24218
|
395 |
|
haftmann@24218
|
396 |
|
haftmann@24589
|
397 |
(** classes and class target **)
|
haftmann@24218
|
398 |
|
haftmann@25002
|
399 |
(* class context syntax *)
|
haftmann@24748
|
400 |
|
haftmann@25083
|
401 |
structure ClassSyntax = ProofDataFun(
|
haftmann@25083
|
402 |
type T = {
|
haftmann@25368
|
403 |
local_constraints: (string * typ) list,
|
haftmann@25368
|
404 |
global_constraints: (string * typ) list,
|
haftmann@25083
|
405 |
base_sort: sort,
|
haftmann@25368
|
406 |
operations: (string * (typ * term)) list,
|
haftmann@25195
|
407 |
unchecks: (term * term) list,
|
haftmann@25083
|
408 |
passed: bool
|
haftmann@25368
|
409 |
};
|
haftmann@25368
|
410 |
fun init _ = {
|
haftmann@25368
|
411 |
local_constraints = [],
|
haftmann@25368
|
412 |
global_constraints = [],
|
haftmann@25368
|
413 |
base_sort = [],
|
haftmann@25368
|
414 |
operations = [],
|
haftmann@25368
|
415 |
unchecks = [],
|
haftmann@25368
|
416 |
passed = true
|
haftmann@25368
|
417 |
};;
|
haftmann@25083
|
418 |
);
|
haftmann@25083
|
419 |
|
wenzelm@25344
|
420 |
fun synchronize_syntax sups base_sort ctxt =
|
haftmann@24914
|
421 |
let
|
wenzelm@25344
|
422 |
val thy = ProofContext.theory_of ctxt;
|
haftmann@25368
|
423 |
fun subst_class_typ sort = map_atyps
|
haftmann@25368
|
424 |
(fn TFree _ => TVar ((Name.aT, 0), sort) | ty' => ty');
|
haftmann@25083
|
425 |
val operations = these_operations thy sups;
|
haftmann@25368
|
426 |
val local_constraints =
|
haftmann@25368
|
427 |
(map o apsnd) (subst_class_typ base_sort o fst o snd) operations;
|
haftmann@25368
|
428 |
val global_constraints =
|
haftmann@25368
|
429 |
(map o apsnd) (fn (class, (ty, _)) => subst_class_typ [class] ty) operations;
|
wenzelm@25318
|
430 |
fun declare_const (c, _) =
|
wenzelm@25318
|
431 |
let val b = Sign.base_name c
|
wenzelm@25344
|
432 |
in Sign.intern_const thy b = c ? Variable.declare_const (b, c) end;
|
haftmann@25368
|
433 |
val unchecks = map (fn (c, (_, (ty, t))) => (t, Const (c, ty))) operations;
|
haftmann@25083
|
434 |
in
|
haftmann@25083
|
435 |
ctxt
|
haftmann@25368
|
436 |
|> fold declare_const local_constraints
|
haftmann@25368
|
437 |
|> fold (ProofContext.add_const_constraint o apsnd SOME) local_constraints
|
haftmann@25368
|
438 |
|> ClassSyntax.put {
|
haftmann@25368
|
439 |
local_constraints = local_constraints,
|
haftmann@25368
|
440 |
global_constraints = global_constraints,
|
haftmann@25083
|
441 |
base_sort = base_sort,
|
haftmann@25368
|
442 |
operations = (map o apsnd) snd operations,
|
haftmann@25195
|
443 |
unchecks = unchecks,
|
haftmann@25083
|
444 |
passed = false
|
haftmann@25368
|
445 |
}
|
haftmann@25083
|
446 |
end;
|
haftmann@25083
|
447 |
|
haftmann@25083
|
448 |
fun refresh_syntax class ctxt =
|
haftmann@25002
|
449 |
let
|
haftmann@25002
|
450 |
val thy = ProofContext.theory_of ctxt;
|
haftmann@25062
|
451 |
val base_sort = (#base_sort o the_class_data thy) class;
|
wenzelm@25344
|
452 |
in synchronize_syntax [class] base_sort ctxt end;
|
haftmann@24914
|
453 |
|
haftmann@25368
|
454 |
val mark_passed = ClassSyntax.map
|
haftmann@25368
|
455 |
(fn { local_constraints, global_constraints, base_sort, operations, unchecks, passed } =>
|
haftmann@25368
|
456 |
{ local_constraints = local_constraints, global_constraints = global_constraints,
|
haftmann@25368
|
457 |
base_sort = base_sort, operations = operations, unchecks = unchecks, passed = true });
|
haftmann@25083
|
458 |
|
haftmann@25083
|
459 |
fun sort_term_check ts ctxt =
|
haftmann@24748
|
460 |
let
|
haftmann@25368
|
461 |
val { local_constraints, global_constraints, base_sort, operations, passed, ... } =
|
haftmann@25368
|
462 |
ClassSyntax.get ctxt;
|
haftmann@25368
|
463 |
fun check_improve (Const (c, ty)) = (case AList.lookup (op =) local_constraints c
|
haftmann@25368
|
464 |
of SOME ty0 => (case try (Type.raw_match (ty0, ty)) Vartab.empty
|
haftmann@25368
|
465 |
of SOME tyenv => (case Vartab.lookup tyenv (Name.aT, 0)
|
haftmann@25368
|
466 |
of SOME (_, TVar (tvar as (vi, _))) =>
|
haftmann@25368
|
467 |
if TypeInfer.is_param vi then cons tvar else I
|
haftmann@25368
|
468 |
| _ => I)
|
haftmann@25368
|
469 |
| NONE => I)
|
haftmann@25368
|
470 |
| NONE => I)
|
haftmann@25368
|
471 |
| check_improve _ = I;
|
haftmann@25368
|
472 |
val improvements = (fold o fold_aterms) check_improve ts [];
|
haftmann@25368
|
473 |
val ts' = (map o map_types o map_atyps) (fn ty as TVar tvar =>
|
haftmann@25368
|
474 |
if member (op =) improvements tvar
|
haftmann@25368
|
475 |
then TFree (Name.aT, base_sort) else ty | ty => ty) ts;
|
haftmann@25368
|
476 |
fun check t0 = Envir.expand_term (fn Const (c, ty) => (case AList.lookup (op =) operations c
|
haftmann@25368
|
477 |
of SOME (ty0, t) =>
|
haftmann@25368
|
478 |
if Type.typ_instance (ProofContext.tsig_of ctxt) (ty, ty0)
|
haftmann@25368
|
479 |
then SOME (ty0, check t) else NONE
|
haftmann@25368
|
480 |
| NONE => NONE)
|
haftmann@25368
|
481 |
| _ => NONE) t0;
|
haftmann@25368
|
482 |
val ts'' = map check ts';
|
haftmann@25368
|
483 |
in if eq_list (op aconv) (ts, ts'') andalso passed then NONE
|
haftmann@25083
|
484 |
else
|
haftmann@25083
|
485 |
ctxt
|
haftmann@25368
|
486 |
|> fold (ProofContext.add_const_constraint o apsnd SOME) global_constraints
|
haftmann@25083
|
487 |
|> mark_passed
|
haftmann@25368
|
488 |
|> pair ts''
|
haftmann@25083
|
489 |
|> SOME
|
haftmann@25083
|
490 |
end;
|
haftmann@24748
|
491 |
|
haftmann@25083
|
492 |
fun sort_term_uncheck ts ctxt =
|
haftmann@25002
|
493 |
let
|
haftmann@25002
|
494 |
val thy = ProofContext.theory_of ctxt;
|
haftmann@25368
|
495 |
val unchecks = (#unchecks o ClassSyntax.get) ctxt;
|
haftmann@25462
|
496 |
val ts' = map (Pattern.rewrite_term thy unchecks []) ts;
|
wenzelm@25060
|
497 |
in if eq_list (op aconv) (ts, ts') then NONE else SOME (ts', ctxt) end;
|
haftmann@25002
|
498 |
|
wenzelm@25344
|
499 |
fun init_ctxt sups base_sort ctxt =
|
haftmann@25083
|
500 |
ctxt
|
haftmann@25083
|
501 |
|> Variable.declare_term
|
haftmann@25083
|
502 |
(Logic.mk_type (TFree (Name.aT, base_sort)))
|
wenzelm@25344
|
503 |
|> synchronize_syntax sups base_sort
|
haftmann@25083
|
504 |
|> Context.proof_map (
|
haftmann@25083
|
505 |
Syntax.add_term_check 0 "class" sort_term_check
|
haftmann@25103
|
506 |
#> Syntax.add_term_uncheck 0 "class" sort_term_uncheck)
|
haftmann@24901
|
507 |
|
haftmann@25311
|
508 |
fun init class thy =
|
haftmann@25311
|
509 |
thy
|
haftmann@25311
|
510 |
|> Locale.init class
|
wenzelm@25344
|
511 |
|> init_ctxt [class] ((#base_sort o the_class_data thy) class);
|
haftmann@24914
|
512 |
|
haftmann@24748
|
513 |
|
haftmann@24589
|
514 |
(* class definition *)
|
haftmann@24218
|
515 |
|
haftmann@24218
|
516 |
local
|
haftmann@24218
|
517 |
|
haftmann@24748
|
518 |
fun gen_class_spec prep_class prep_expr process_expr thy raw_supclasses raw_includes_elems =
|
haftmann@24218
|
519 |
let
|
haftmann@24748
|
520 |
val supclasses = map (prep_class thy) raw_supclasses;
|
haftmann@24748
|
521 |
val supsort = Sign.minimize_sort thy supclasses;
|
haftmann@25618
|
522 |
val sups = filter (is_class thy) supsort;
|
haftmann@25618
|
523 |
val base_sort = if null sups then supsort else
|
haftmann@25618
|
524 |
(#base_sort o the_class_data thy o hd) sups;
|
haftmann@25038
|
525 |
val suplocales = map Locale.Locale sups;
|
haftmann@24748
|
526 |
val (raw_elems, includes) = fold_rev (fn Locale.Elem e => apfst (cons e)
|
haftmann@24748
|
527 |
| Locale.Expr i => apsnd (cons (prep_expr thy i))) raw_includes_elems ([], []);
|
haftmann@24748
|
528 |
val supexpr = Locale.Merge suplocales;
|
haftmann@24748
|
529 |
val supparams = (map fst o Locale.parameters_of_expr thy) supexpr;
|
haftmann@24748
|
530 |
val mergeexpr = Locale.Merge (suplocales @ includes);
|
haftmann@24748
|
531 |
val constrain = Element.Constrains ((map o apsnd o map_atyps)
|
wenzelm@24847
|
532 |
(fn TFree (_, sort) => TFree (Name.aT, sort)) supparams);
|
haftmann@25683
|
533 |
fun fork_syn (Element.Fixes xs) =
|
haftmann@25683
|
534 |
fold_map (fn (c, ty, syn) => cons (c, syn) #> pair (c, ty, NoSyn)) xs
|
haftmann@25683
|
535 |
#>> Element.Fixes
|
haftmann@25683
|
536 |
| fork_syn x = pair x;
|
haftmann@25683
|
537 |
fun fork_syntax elems =
|
haftmann@25683
|
538 |
let
|
haftmann@25683
|
539 |
val (elems', global_syntax) = fold_map fork_syn elems [];
|
haftmann@25683
|
540 |
in (if null includes (*FIXME*) then constrain :: elems' else elems', global_syntax) end;
|
haftmann@25683
|
541 |
val (elems, global_syntax) =
|
haftmann@25683
|
542 |
ProofContext.init thy
|
haftmann@25683
|
543 |
|> Locale.cert_expr supexpr [constrain]
|
haftmann@25683
|
544 |
|> snd
|
haftmann@25683
|
545 |
|> init_ctxt sups base_sort
|
haftmann@25683
|
546 |
|> process_expr Locale.empty raw_elems
|
haftmann@25683
|
547 |
|> fst
|
haftmann@25683
|
548 |
|> fork_syntax
|
haftmann@25683
|
549 |
in (((sups, supparams), (supsort, base_sort, mergeexpr)), (elems, global_syntax)) end;
|
haftmann@24748
|
550 |
|
haftmann@24748
|
551 |
val read_class_spec = gen_class_spec Sign.intern_class Locale.intern_expr Locale.read_expr;
|
haftmann@24748
|
552 |
val check_class_spec = gen_class_spec (K I) (K I) Locale.cert_expr;
|
haftmann@24748
|
553 |
|
haftmann@25683
|
554 |
fun adjungate_axclass bname class base_sort sups supsort supparams global_syntax other_consts thy =
|
wenzelm@24968
|
555 |
let
|
haftmann@25683
|
556 |
val supconsts = map fst supparams
|
haftmann@25683
|
557 |
|> AList.make (the o AList.lookup (op =) (these_params thy sups))
|
haftmann@25683
|
558 |
|> (map o apsnd o apsnd o map_atyps o K o TFree) (Name.aT, [class]);
|
haftmann@25683
|
559 |
val all_params = map fst (Locale.parameters_of thy class);
|
haftmann@25683
|
560 |
fun add_const (v, raw_ty) thy =
|
haftmann@25683
|
561 |
let
|
haftmann@25683
|
562 |
val c = Sign.full_name thy v;
|
haftmann@25683
|
563 |
val ty = map_atyps (K (TFree (Name.aT, base_sort))) raw_ty;
|
haftmann@25683
|
564 |
val ty0 = Type.strip_sorts ty;
|
haftmann@25683
|
565 |
val ty' = map_atyps (K (TFree (Name.aT, [class]))) ty0;
|
haftmann@25683
|
566 |
val syn = (the_default NoSyn o AList.lookup (op =) global_syntax) v;
|
haftmann@25683
|
567 |
in
|
haftmann@25683
|
568 |
thy
|
haftmann@25683
|
569 |
|> Sign.declare_const [] (v, ty0, syn)
|
haftmann@25683
|
570 |
|> snd
|
haftmann@25683
|
571 |
|> pair ((v, ty), (c, ty'))
|
haftmann@25683
|
572 |
end;
|
haftmann@25683
|
573 |
fun add_consts raw_params thy =
|
haftmann@25683
|
574 |
thy
|
haftmann@25683
|
575 |
|> Sign.add_path (Logic.const_of_class bname)
|
haftmann@25683
|
576 |
|> fold_map add_const raw_params
|
haftmann@25683
|
577 |
||> Sign.restore_naming thy
|
haftmann@25683
|
578 |
|-> (fn params => pair (supconsts @ (map o apfst) fst params, params));
|
haftmann@25683
|
579 |
fun globalize param_map = map_aterms
|
haftmann@25683
|
580 |
(fn Free (v, ty) => Const ((fst o the o AList.lookup (op =) param_map) v, ty)
|
haftmann@25683
|
581 |
| t => t);
|
haftmann@25683
|
582 |
val raw_pred = Locale.intros thy class
|
haftmann@25683
|
583 |
|> fst
|
haftmann@25683
|
584 |
|> map (Logic.unvarify o Logic.strip_imp_concl o Thm.prop_of);
|
haftmann@25683
|
585 |
fun get_axiom thy = case (#axioms o AxClass.get_info thy) class
|
haftmann@25683
|
586 |
of [] => NONE
|
haftmann@25683
|
587 |
| [thm] => SOME thm;
|
wenzelm@24968
|
588 |
in
|
wenzelm@24968
|
589 |
thy
|
haftmann@25683
|
590 |
|> add_consts ((snd o chop (length supparams)) all_params)
|
haftmann@25683
|
591 |
|-> (fn (param_map, params) => AxClass.define_class (bname, supsort)
|
haftmann@25683
|
592 |
(map (fst o snd) params @ other_consts)
|
haftmann@25683
|
593 |
[((bname ^ "_" ^ AxClass.axiomsN, []), map (globalize param_map) raw_pred)]
|
haftmann@25683
|
594 |
#> snd
|
haftmann@25683
|
595 |
#> `get_axiom
|
haftmann@25683
|
596 |
#-> (fn assm_axiom => fold (Sign.add_const_constraint o apsnd SOME o snd) params
|
haftmann@25683
|
597 |
#> pair (param_map, params, assm_axiom)))
|
wenzelm@24968
|
598 |
end;
|
wenzelm@24968
|
599 |
|
haftmann@25002
|
600 |
fun gen_class prep_spec prep_param bname
|
haftmann@24748
|
601 |
raw_supclasses raw_includes_elems raw_other_consts thy =
|
haftmann@24748
|
602 |
let
|
haftmann@25038
|
603 |
val class = Sign.full_name thy bname;
|
haftmann@25683
|
604 |
val (((sups, supparams), (supsort, base_sort, mergeexpr)), (elems, global_syntax)) =
|
haftmann@24748
|
605 |
prep_spec thy raw_supclasses raw_includes_elems;
|
wenzelm@24968
|
606 |
val other_consts = map (tap (Sign.the_const_type thy) o prep_param thy) raw_other_consts;
|
haftmann@24218
|
607 |
in
|
haftmann@24218
|
608 |
thy
|
haftmann@24748
|
609 |
|> Locale.add_locale_i (SOME "") bname mergeexpr elems
|
haftmann@25038
|
610 |
|> snd
|
haftmann@25311
|
611 |
|> ProofContext.theory_of
|
haftmann@25683
|
612 |
|> adjungate_axclass bname class base_sort sups supsort supparams global_syntax other_consts
|
haftmann@25683
|
613 |
|-> (fn (param_map, params, assm_axiom) =>
|
haftmann@25683
|
614 |
`(fn thy => calculate thy sups base_sort assm_axiom param_map class)
|
haftmann@25711
|
615 |
#-> (fn (morphism, axiom, assm_intro, of_class) =>
|
haftmann@25683
|
616 |
add_class_data ((class, sups), (params, base_sort,
|
haftmann@25711
|
617 |
map snd param_map, morphism, axiom, assm_intro, of_class))
|
haftmann@25683
|
618 |
#> class_interpretation class (the_list axiom) []))
|
haftmann@25268
|
619 |
|> init class
|
haftmann@25038
|
620 |
|> pair class
|
haftmann@24218
|
621 |
end;
|
haftmann@24218
|
622 |
|
wenzelm@25326
|
623 |
fun read_const thy = #1 o Term.dest_Const o ProofContext.read_const (ProofContext.init thy);
|
wenzelm@25326
|
624 |
|
haftmann@24218
|
625 |
in
|
haftmann@24218
|
626 |
|
wenzelm@25326
|
627 |
val class_cmd = gen_class read_class_spec read_const;
|
haftmann@24748
|
628 |
val class = gen_class check_class_spec (K I);
|
haftmann@24218
|
629 |
|
haftmann@24218
|
630 |
end; (*local*)
|
haftmann@24218
|
631 |
|
haftmann@24218
|
632 |
|
haftmann@25485
|
633 |
(* class target *)
|
haftmann@24218
|
634 |
|
haftmann@25603
|
635 |
fun declare class pos ((c, mx), dict) thy =
|
haftmann@24218
|
636 |
let
|
wenzelm@25024
|
637 |
val prfx = class_prefix class;
|
wenzelm@25024
|
638 |
val thy' = thy |> Sign.add_path prfx;
|
haftmann@25711
|
639 |
val phi = partial_morphism thy' class;
|
wenzelm@25024
|
640 |
|
haftmann@25062
|
641 |
val c' = Sign.full_name thy' c;
|
haftmann@25239
|
642 |
val dict' = Morphism.term phi dict;
|
haftmann@25239
|
643 |
val dict_def = map_types Logic.unvarifyT dict';
|
haftmann@25239
|
644 |
val ty' = Term.fastype_of dict_def;
|
haftmann@25083
|
645 |
val ty'' = Type.strip_sorts ty';
|
haftmann@25239
|
646 |
val def_eq = Logic.mk_equals (Const (c', ty'), dict_def);
|
haftmann@25618
|
647 |
fun get_axiom thy = ((Thm.varifyT o Thm.symmetric o Thm.get_axiom_i thy) c', thy);
|
haftmann@24218
|
648 |
in
|
wenzelm@25024
|
649 |
thy'
|
haftmann@25096
|
650 |
|> Sign.declare_const pos (c, ty'', mx) |> snd
|
haftmann@25518
|
651 |
|> Thm.add_def false false (c, def_eq)
|
haftmann@25062
|
652 |
|>> Thm.symmetric
|
haftmann@25618
|
653 |
||>> get_axiom
|
haftmann@25618
|
654 |
|-> (fn (def, def') => class_interpretation class [def] [Thm.prop_of def]
|
haftmann@25618
|
655 |
#> register_operation class (c', (dict', SOME def')))
|
haftmann@24218
|
656 |
|> Sign.restore_naming thy
|
haftmann@25083
|
657 |
|> Sign.add_const_constraint (c', SOME ty')
|
haftmann@24218
|
658 |
end;
|
haftmann@24218
|
659 |
|
haftmann@25603
|
660 |
fun abbrev class prmode pos ((c, mx), rhs) thy =
|
haftmann@24836
|
661 |
let
|
wenzelm@25024
|
662 |
val prfx = class_prefix class;
|
haftmann@25096
|
663 |
val thy' = thy |> Sign.add_path prfx;
|
haftmann@25062
|
664 |
val phi = morphism thy class;
|
haftmann@25062
|
665 |
|
haftmann@25096
|
666 |
val c' = Sign.full_name thy' c;
|
haftmann@25711
|
667 |
val rhs' = Morphism.term phi rhs;
|
haftmann@25239
|
668 |
val ty' = Logic.unvarifyT (Term.fastype_of rhs');
|
haftmann@24836
|
669 |
in
|
haftmann@25096
|
670 |
thy'
|
haftmann@25146
|
671 |
|> Sign.add_abbrev (#1 prmode) pos (c, map_types Type.strip_sorts rhs') |> snd
|
haftmann@25096
|
672 |
|> Sign.add_const_constraint (c', SOME ty')
|
wenzelm@25024
|
673 |
|> Sign.notation true prmode [(Const (c', ty'), mx)]
|
haftmann@25368
|
674 |
|> register_operation class (c', (rhs', NONE))
|
haftmann@25096
|
675 |
|> Sign.restore_naming thy
|
haftmann@24836
|
676 |
end;
|
haftmann@24836
|
677 |
|
haftmann@25462
|
678 |
|
haftmann@25462
|
679 |
(** instantiation target **)
|
haftmann@25462
|
680 |
|
haftmann@25462
|
681 |
(* bookkeeping *)
|
haftmann@25462
|
682 |
|
haftmann@25462
|
683 |
datatype instantiation = Instantiation of {
|
haftmann@25864
|
684 |
arities: string list * (string * sort) list * sort,
|
haftmann@25462
|
685 |
params: ((string * string) * (string * typ)) list
|
haftmann@25603
|
686 |
(*(instantiation parameter, type constructor), (local instantiation parameter, typ)*)
|
haftmann@25462
|
687 |
}
|
haftmann@25462
|
688 |
|
haftmann@25462
|
689 |
structure Instantiation = ProofDataFun
|
haftmann@25462
|
690 |
(
|
haftmann@25462
|
691 |
type T = instantiation
|
haftmann@25536
|
692 |
fun init _ = Instantiation { arities = ([], [], []), params = [] };
|
haftmann@25462
|
693 |
);
|
haftmann@25462
|
694 |
|
haftmann@25485
|
695 |
fun mk_instantiation (arities, params) =
|
haftmann@25485
|
696 |
Instantiation { arities = arities, params = params };
|
haftmann@25514
|
697 |
fun get_instantiation lthy = case Instantiation.get (LocalTheory.target_of lthy)
|
haftmann@25485
|
698 |
of Instantiation data => data;
|
haftmann@25514
|
699 |
fun map_instantiation f = (LocalTheory.target o Instantiation.map)
|
haftmann@25514
|
700 |
(fn Instantiation { arities, params } => mk_instantiation (f (arities, params)));
|
haftmann@25462
|
701 |
|
haftmann@25514
|
702 |
fun the_instantiation lthy = case get_instantiation lthy
|
haftmann@25536
|
703 |
of { arities = ([], [], []), ... } => error "No instantiation target"
|
haftmann@25485
|
704 |
| data => data;
|
haftmann@25462
|
705 |
|
haftmann@25485
|
706 |
val instantiation_params = #params o get_instantiation;
|
haftmann@25462
|
707 |
|
haftmann@25514
|
708 |
fun instantiation_param lthy v = instantiation_params lthy
|
haftmann@25462
|
709 |
|> find_first (fn (_, (v', _)) => v = v')
|
haftmann@25462
|
710 |
|> Option.map (fst o fst);
|
haftmann@25462
|
711 |
|
haftmann@25514
|
712 |
fun confirm_declaration c = (map_instantiation o apsnd)
|
haftmann@25485
|
713 |
(filter_out (fn (_, (c', _)) => c' = c));
|
haftmann@25485
|
714 |
|
haftmann@25462
|
715 |
|
haftmann@25462
|
716 |
(* syntax *)
|
haftmann@25462
|
717 |
|
haftmann@25514
|
718 |
fun inst_term_check ts lthy =
|
haftmann@25462
|
719 |
let
|
haftmann@25514
|
720 |
val params = instantiation_params lthy;
|
haftmann@25514
|
721 |
val tsig = ProofContext.tsig_of lthy;
|
haftmann@25514
|
722 |
val thy = ProofContext.theory_of lthy;
|
haftmann@25462
|
723 |
|
haftmann@25597
|
724 |
fun check_improve (Const (c, ty)) = (case AxClass.inst_tyco_of thy (c, ty)
|
haftmann@25462
|
725 |
of SOME tyco => (case AList.lookup (op =) params (c, tyco)
|
haftmann@25502
|
726 |
of SOME (_, ty') => perhaps (try (Type.typ_match tsig (ty, ty')))
|
haftmann@25462
|
727 |
| NONE => I)
|
haftmann@25462
|
728 |
| NONE => I)
|
haftmann@25462
|
729 |
| check_improve _ = I;
|
haftmann@25711
|
730 |
val subst_param = map_aterms (fn t as Const (c, ty) =>
|
haftmann@25711
|
731 |
(case AxClass.inst_tyco_of thy (c, ty)
|
haftmann@25711
|
732 |
of SOME tyco => (case AList.lookup (op =) params (c, tyco)
|
haftmann@25711
|
733 |
of SOME v_ty => Free v_ty
|
haftmann@25711
|
734 |
| NONE => t)
|
haftmann@25711
|
735 |
| NONE => t)
|
haftmann@25711
|
736 |
| t => t);
|
haftmann@25711
|
737 |
|
haftmann@25462
|
738 |
val improvement = (fold o fold_aterms) check_improve ts Vartab.empty;
|
haftmann@25462
|
739 |
val ts' = (map o map_types) (Envir.typ_subst_TVars improvement) ts;
|
haftmann@25711
|
740 |
val ts'' = map subst_param ts';
|
haftmann@25514
|
741 |
in if eq_list (op aconv) (ts, ts'') then NONE else SOME (ts'', lthy) end;
|
haftmann@25462
|
742 |
|
haftmann@25514
|
743 |
fun inst_term_uncheck ts lthy =
|
haftmann@25462
|
744 |
let
|
haftmann@25514
|
745 |
val params = instantiation_params lthy;
|
haftmann@25462
|
746 |
val ts' = (map o map_aterms) (fn t as Free (v, ty) =>
|
haftmann@25462
|
747 |
(case get_first (fn ((c, _), (v', _)) => if v = v' then SOME c else NONE) params
|
haftmann@25462
|
748 |
of SOME c => Const (c, ty)
|
haftmann@25462
|
749 |
| NONE => t)
|
haftmann@25462
|
750 |
| t => t) ts;
|
haftmann@25514
|
751 |
in if eq_list (op aconv) (ts, ts') then NONE else SOME (ts', lthy) end;
|
haftmann@25462
|
752 |
|
haftmann@25462
|
753 |
|
haftmann@25462
|
754 |
(* target *)
|
haftmann@25462
|
755 |
|
haftmann@25485
|
756 |
val sanatize_name = (*FIXME*)
|
haftmann@25485
|
757 |
let
|
haftmann@25574
|
758 |
fun is_valid s = Symbol.is_ascii_letter s orelse Symbol.is_ascii_digit s
|
haftmann@25574
|
759 |
orelse s = "'" orelse s = "_";
|
haftmann@25485
|
760 |
val is_junk = not o is_valid andf Symbol.is_regular;
|
haftmann@25485
|
761 |
val junk = Scan.many is_junk;
|
haftmann@25485
|
762 |
val scan_valids = Symbol.scanner "Malformed input"
|
haftmann@25485
|
763 |
((junk |--
|
haftmann@25485
|
764 |
(Scan.optional (Scan.one Symbol.is_ascii_letter) "x" ^^ (Scan.many is_valid >> implode)
|
haftmann@25485
|
765 |
--| junk))
|
wenzelm@25999
|
766 |
::: Scan.repeat ((Scan.many1 is_valid >> implode) --| junk));
|
haftmann@25485
|
767 |
in
|
haftmann@25485
|
768 |
explode #> scan_valids #> implode
|
haftmann@25485
|
769 |
end;
|
haftmann@25485
|
770 |
|
haftmann@25864
|
771 |
fun init_instantiation (tycos, vs, sort) thy =
|
haftmann@25462
|
772 |
let
|
haftmann@25536
|
773 |
val _ = if null tycos then error "At least one arity must be given" else ();
|
haftmann@25536
|
774 |
val _ = map (the_class_data thy) sort;
|
haftmann@25485
|
775 |
fun type_name "*" = "prod"
|
haftmann@25485
|
776 |
| type_name "+" = "sum"
|
haftmann@25485
|
777 |
| type_name s = sanatize_name (NameSpace.base s); (*FIXME*)
|
haftmann@25597
|
778 |
fun get_param tyco (param, (c, ty)) = if can (AxClass.param_of_inst thy) (c, tyco)
|
haftmann@25603
|
779 |
then NONE else SOME ((c, tyco),
|
haftmann@25864
|
780 |
(param ^ "_" ^ type_name tyco, map_atyps (K (Type (tyco, map TFree vs))) ty));
|
haftmann@25536
|
781 |
val params = map_product get_param tycos (these_params thy sort) |> map_filter I;
|
haftmann@25485
|
782 |
in
|
haftmann@25485
|
783 |
thy
|
haftmann@25485
|
784 |
|> ProofContext.init
|
haftmann@25864
|
785 |
|> Instantiation.put (mk_instantiation ((tycos, vs, sort), params))
|
haftmann@25864
|
786 |
|> fold (Variable.declare_term o Logic.mk_type o TFree) vs
|
haftmann@25574
|
787 |
|> fold (Variable.declare_names o Free o snd) params
|
haftmann@25864
|
788 |
|> fold (fn tyco => ProofContext.add_arity (tyco, map snd vs, sort)) tycos
|
haftmann@25485
|
789 |
|> Context.proof_map (
|
haftmann@25485
|
790 |
Syntax.add_term_check 0 "instance" inst_term_check
|
haftmann@25485
|
791 |
#> Syntax.add_term_uncheck 0 "instance" inst_term_uncheck)
|
haftmann@25485
|
792 |
end;
|
haftmann@25485
|
793 |
|
haftmann@25485
|
794 |
fun gen_instantiation_instance do_proof after_qed lthy =
|
haftmann@25485
|
795 |
let
|
haftmann@25864
|
796 |
val (tycos, vs, sort) = (#arities o the_instantiation) lthy;
|
haftmann@25864
|
797 |
val arities_proof = maps (fn tyco => Logic.mk_arities (tyco, map snd vs, sort)) tycos;
|
haftmann@25462
|
798 |
fun after_qed' results =
|
haftmann@25462
|
799 |
LocalTheory.theory (fold (AxClass.add_arity o Thm.varifyT) results)
|
haftmann@25462
|
800 |
#> after_qed;
|
haftmann@25462
|
801 |
in
|
haftmann@25462
|
802 |
lthy
|
haftmann@25462
|
803 |
|> do_proof after_qed' arities_proof
|
haftmann@25462
|
804 |
end;
|
haftmann@25462
|
805 |
|
haftmann@25485
|
806 |
val instantiation_instance = gen_instantiation_instance (fn after_qed => fn ts =>
|
haftmann@25462
|
807 |
Proof.theorem_i NONE (after_qed o map the_single) (map (fn t => [(t, [])]) ts));
|
haftmann@25462
|
808 |
|
haftmann@25485
|
809 |
fun prove_instantiation_instance tac = gen_instantiation_instance (fn after_qed =>
|
haftmann@25502
|
810 |
fn ts => fn lthy => after_qed (map (fn t => Goal.prove lthy [] [] t
|
haftmann@25502
|
811 |
(fn {context, ...} => tac context)) ts) lthy) I;
|
haftmann@25462
|
812 |
|
haftmann@25462
|
813 |
fun conclude_instantiation lthy =
|
haftmann@25462
|
814 |
let
|
haftmann@25485
|
815 |
val { arities, params } = the_instantiation lthy;
|
haftmann@25864
|
816 |
val (tycos, vs, sort) = arities;
|
haftmann@25462
|
817 |
val thy = ProofContext.theory_of lthy;
|
haftmann@25597
|
818 |
val _ = map (fn tyco => if Sign.of_sort thy
|
haftmann@25864
|
819 |
(Type (tyco, map TFree vs), sort)
|
haftmann@25462
|
820 |
then () else error ("Missing instance proof for type " ^ quote (Sign.extern_type thy tyco)))
|
haftmann@25597
|
821 |
tycos;
|
haftmann@25770
|
822 |
(*this checkpoint should move to AxClass as soon as "attach" has disappeared*)
|
haftmann@25770
|
823 |
val _ = case map (fst o snd) params
|
haftmann@25770
|
824 |
of [] => ()
|
haftmann@25770
|
825 |
| cs => Output.legacy_feature
|
haftmann@25829
|
826 |
("Missing specifications for overloaded parameters " ^ commas_quote cs)
|
haftmann@25597
|
827 |
in lthy end;
|
haftmann@25462
|
828 |
|
haftmann@25603
|
829 |
fun pretty_instantiation lthy =
|
haftmann@25603
|
830 |
let
|
haftmann@25603
|
831 |
val { arities, params } = the_instantiation lthy;
|
haftmann@25864
|
832 |
val (tycos, vs, sort) = arities;
|
haftmann@25603
|
833 |
val thy = ProofContext.theory_of lthy;
|
haftmann@25864
|
834 |
fun pr_arity tyco = Syntax.pretty_arity lthy (tyco, map snd vs, sort);
|
haftmann@25603
|
835 |
fun pr_param ((c, _), (v, ty)) =
|
haftmann@25864
|
836 |
(Pretty.block o Pretty.breaks) [Pretty.str v, Pretty.str "==",
|
haftmann@25864
|
837 |
(Pretty.str o Sign.extern_const thy) c, Pretty.str "::", Sign.pretty_typ thy ty];
|
haftmann@25603
|
838 |
in
|
haftmann@25603
|
839 |
(Pretty.block o Pretty.fbreaks)
|
haftmann@25603
|
840 |
(Pretty.str "instantiation" :: map pr_arity tycos @ map pr_param params)
|
haftmann@25603
|
841 |
end;
|
haftmann@25603
|
842 |
|
haftmann@24218
|
843 |
end;
|
haftmann@25683
|
844 |
|