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