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@25038
|
15 |
val is_class: theory -> class -> bool
|
haftmann@25002
|
16 |
val these_params: theory -> sort -> (string * (string * typ)) list
|
haftmann@25311
|
17 |
val init: class -> theory -> Proof.context
|
haftmann@25096
|
18 |
val add_logical_const: string -> Markup.property list
|
wenzelm@25104
|
19 |
-> (string * mixfix) * term -> theory -> theory
|
haftmann@25096
|
20 |
val add_syntactic_const: string -> Syntax.mode -> Markup.property list
|
wenzelm@25104
|
21 |
-> (string * mixfix) * term -> theory -> theory
|
haftmann@25083
|
22 |
val refresh_syntax: class -> Proof.context -> Proof.context
|
haftmann@24589
|
23 |
val intro_classes_tac: thm list -> tactic
|
haftmann@24589
|
24 |
val default_intro_classes_tac: thm list -> tactic
|
haftmann@25195
|
25 |
val prove_subclass: class * class -> thm list -> Proof.context
|
haftmann@25195
|
26 |
-> theory -> theory
|
haftmann@24589
|
27 |
val print_classes: theory -> unit
|
wenzelm@25210
|
28 |
val class_prefix: string -> string
|
haftmann@24423
|
29 |
|
haftmann@25462
|
30 |
(*instances*)
|
haftmann@25462
|
31 |
val declare_overloaded: string * typ * mixfix -> theory -> term * theory
|
haftmann@25462
|
32 |
val define_overloaded: string -> string * term -> theory -> thm * theory
|
haftmann@25462
|
33 |
val unoverload: theory -> conv
|
haftmann@25462
|
34 |
val overload: theory -> conv
|
haftmann@25462
|
35 |
val unoverload_const: theory -> string * typ -> string
|
haftmann@25462
|
36 |
val inst_const: theory -> string * string -> string
|
haftmann@25462
|
37 |
val param_const: theory -> string -> (string * string) option
|
haftmann@25462
|
38 |
val instantiation: arity list -> theory -> local_theory
|
haftmann@25462
|
39 |
val proof_instantiation: (local_theory -> local_theory) -> local_theory -> Proof.state
|
haftmann@25462
|
40 |
val prove_instantiation: (Proof.context -> tactic) -> local_theory -> local_theory
|
haftmann@25462
|
41 |
val conclude_instantiation: local_theory -> local_theory
|
haftmann@25462
|
42 |
val end_instantiation: local_theory -> Proof.context
|
haftmann@25462
|
43 |
val instantiation_const: Proof.context -> string -> string option
|
haftmann@25462
|
44 |
|
haftmann@25462
|
45 |
(*old axclass layer*)
|
haftmann@25462
|
46 |
val axclass_cmd: bstring * xstring list
|
haftmann@25462
|
47 |
-> ((bstring * Attrib.src list) * string list) list
|
haftmann@25462
|
48 |
-> theory -> class * theory
|
haftmann@25462
|
49 |
val classrel_cmd: xstring * xstring -> theory -> Proof.state
|
haftmann@25462
|
50 |
|
haftmann@25462
|
51 |
(*old instance layer*)
|
haftmann@24589
|
52 |
val instance_arity: (theory -> theory) -> arity list -> theory -> Proof.state
|
haftmann@24589
|
53 |
val instance: arity list -> ((bstring * Attrib.src list) * term) list
|
haftmann@24423
|
54 |
-> (thm list -> theory -> theory)
|
haftmann@24218
|
55 |
-> theory -> Proof.state
|
haftmann@24589
|
56 |
val instance_cmd: (bstring * xstring list * xstring) list
|
haftmann@24589
|
57 |
-> ((bstring * Attrib.src list) * xstring) list
|
haftmann@24423
|
58 |
-> (thm list -> theory -> theory)
|
haftmann@24218
|
59 |
-> theory -> Proof.state
|
haftmann@24589
|
60 |
val prove_instance: tactic -> arity list
|
haftmann@24218
|
61 |
-> ((bstring * Attrib.src list) * term) list
|
haftmann@24423
|
62 |
-> theory -> thm list * theory
|
haftmann@24218
|
63 |
end;
|
haftmann@24218
|
64 |
|
haftmann@24218
|
65 |
structure Class : CLASS =
|
haftmann@24218
|
66 |
struct
|
haftmann@24218
|
67 |
|
haftmann@24218
|
68 |
(** auxiliary **)
|
haftmann@24218
|
69 |
|
haftmann@25062
|
70 |
val classN = "class";
|
haftmann@25062
|
71 |
val introN = "intro";
|
haftmann@25062
|
72 |
|
haftmann@25002
|
73 |
fun prove_interpretation tac prfx_atts expr inst =
|
haftmann@25002
|
74 |
Locale.interpretation_i I prfx_atts expr inst
|
haftmann@24589
|
75 |
#> Proof.global_terminal_proof
|
haftmann@24589
|
76 |
(Method.Basic (K (Method.SIMPLE_METHOD tac), Position.none), NONE)
|
haftmann@24589
|
77 |
#> ProofContext.theory_of;
|
haftmann@24589
|
78 |
|
haftmann@25195
|
79 |
fun prove_interpretation_in tac after_qed (name, expr) =
|
haftmann@25195
|
80 |
Locale.interpretation_in_locale
|
haftmann@25195
|
81 |
(ProofContext.theory after_qed) (name, expr)
|
haftmann@25195
|
82 |
#> Proof.global_terminal_proof
|
haftmann@25195
|
83 |
(Method.Basic (K (Method.SIMPLE_METHOD tac), Position.none), NONE)
|
haftmann@25195
|
84 |
#> ProofContext.theory_of;
|
haftmann@25195
|
85 |
|
wenzelm@25020
|
86 |
fun OF_LAST thm1 thm2 = thm1 RSN (Thm.nprems_of thm2, thm2);
|
haftmann@24589
|
87 |
|
haftmann@24589
|
88 |
fun strip_all_ofclass thy sort =
|
haftmann@24589
|
89 |
let
|
wenzelm@24847
|
90 |
val typ = TVar ((Name.aT, 0), sort);
|
haftmann@24589
|
91 |
fun prem_inclass t =
|
haftmann@24589
|
92 |
case Logic.strip_imp_prems t
|
haftmann@24589
|
93 |
of ofcls :: _ => try Logic.dest_inclass ofcls
|
haftmann@24589
|
94 |
| [] => NONE;
|
haftmann@24589
|
95 |
fun strip_ofclass class thm =
|
haftmann@24589
|
96 |
thm OF (fst o AxClass.of_sort thy (typ, [class])) AxClass.cache;
|
haftmann@24589
|
97 |
fun strip thm = case (prem_inclass o Thm.prop_of) thm
|
haftmann@24589
|
98 |
of SOME (_, class) => thm |> strip_ofclass class |> strip
|
haftmann@24589
|
99 |
| NONE => thm;
|
haftmann@24589
|
100 |
in strip end;
|
haftmann@24589
|
101 |
|
haftmann@25038
|
102 |
fun get_remove_global_constraint c thy =
|
haftmann@25038
|
103 |
let
|
haftmann@25038
|
104 |
val ty = Sign.the_const_constraint thy c;
|
haftmann@25038
|
105 |
in
|
haftmann@25038
|
106 |
thy
|
haftmann@25038
|
107 |
|> Sign.add_const_constraint (c, NONE)
|
haftmann@25038
|
108 |
|> pair (c, Logic.unvarifyT ty)
|
haftmann@25038
|
109 |
end;
|
haftmann@25038
|
110 |
|
haftmann@24589
|
111 |
|
haftmann@24589
|
112 |
(** axclass command **)
|
haftmann@24589
|
113 |
|
haftmann@24218
|
114 |
fun axclass_cmd (class, raw_superclasses) raw_specs thy =
|
haftmann@24218
|
115 |
let
|
haftmann@24218
|
116 |
val ctxt = ProofContext.init thy;
|
haftmann@24218
|
117 |
val superclasses = map (Sign.read_class thy) raw_superclasses;
|
haftmann@24589
|
118 |
val name_atts = map ((apsnd o map) (Attrib.attribute thy) o fst)
|
haftmann@24589
|
119 |
raw_specs;
|
haftmann@24589
|
120 |
val axiomss = ProofContext.read_propp (ctxt, map (map (rpair []) o snd)
|
haftmann@24589
|
121 |
raw_specs)
|
haftmann@24218
|
122 |
|> snd
|
haftmann@24218
|
123 |
|> (map o map) fst;
|
haftmann@24589
|
124 |
in
|
haftmann@24589
|
125 |
AxClass.define_class (class, superclasses) []
|
haftmann@24589
|
126 |
(name_atts ~~ axiomss) thy
|
haftmann@24589
|
127 |
end;
|
haftmann@24218
|
128 |
|
haftmann@24218
|
129 |
local
|
haftmann@24218
|
130 |
|
haftmann@24218
|
131 |
fun gen_instance mk_prop add_thm after_qed insts thy =
|
haftmann@24218
|
132 |
let
|
haftmann@24218
|
133 |
fun after_qed' results =
|
haftmann@24218
|
134 |
ProofContext.theory ((fold o fold) add_thm results #> after_qed);
|
haftmann@24218
|
135 |
in
|
haftmann@24218
|
136 |
thy
|
haftmann@24218
|
137 |
|> ProofContext.init
|
haftmann@24589
|
138 |
|> Proof.theorem_i NONE after_qed' ((map (fn t => [(t, [])])
|
haftmann@24589
|
139 |
o maps (mk_prop thy)) insts)
|
haftmann@24218
|
140 |
end;
|
haftmann@24218
|
141 |
|
haftmann@24218
|
142 |
in
|
haftmann@24218
|
143 |
|
haftmann@24589
|
144 |
val instance_arity =
|
haftmann@24218
|
145 |
gen_instance (Logic.mk_arities oo Sign.cert_arity) AxClass.add_arity;
|
haftmann@24589
|
146 |
val classrel =
|
haftmann@24218
|
147 |
gen_instance (single oo (Logic.mk_classrel oo AxClass.cert_classrel))
|
haftmann@24218
|
148 |
AxClass.add_classrel I o single;
|
haftmann@24589
|
149 |
val classrel_cmd =
|
haftmann@24589
|
150 |
gen_instance (single oo (Logic.mk_classrel oo AxClass.read_classrel))
|
haftmann@24589
|
151 |
AxClass.add_classrel I o single;
|
haftmann@24218
|
152 |
|
haftmann@24218
|
153 |
end; (*local*)
|
haftmann@24218
|
154 |
|
haftmann@24218
|
155 |
|
haftmann@25462
|
156 |
(** basic overloading **)
|
haftmann@25462
|
157 |
|
haftmann@25462
|
158 |
(* bookkeeping *)
|
haftmann@24304
|
159 |
|
haftmann@24304
|
160 |
structure InstData = TheoryDataFun
|
haftmann@24304
|
161 |
(
|
haftmann@24423
|
162 |
type T = (string * thm) Symtab.table Symtab.table * (string * string) Symtab.table;
|
haftmann@24423
|
163 |
(*constant name ~> type constructor ~> (constant name, equation),
|
haftmann@24423
|
164 |
constant name ~> (constant name, type constructor)*)
|
haftmann@24423
|
165 |
val empty = (Symtab.empty, Symtab.empty);
|
haftmann@24304
|
166 |
val copy = I;
|
haftmann@24304
|
167 |
val extend = I;
|
haftmann@24423
|
168 |
fun merge _ ((taba1, tabb1), (taba2, tabb2)) =
|
haftmann@24423
|
169 |
(Symtab.join (K (Symtab.merge (K true))) (taba1, taba2),
|
haftmann@24423
|
170 |
Symtab.merge (K true) (tabb1, tabb2));
|
haftmann@24304
|
171 |
);
|
haftmann@24304
|
172 |
|
haftmann@25462
|
173 |
val inst_tyco = Option.map fst o try (dest_Type o the_single) oo Sign.const_typargs;
|
haftmann@25462
|
174 |
|
haftmann@25462
|
175 |
fun inst thy (c, tyco) =
|
haftmann@25462
|
176 |
(the o Symtab.lookup ((the o Symtab.lookup (fst (InstData.get thy))) c)) tyco;
|
haftmann@25462
|
177 |
|
haftmann@25462
|
178 |
val inst_const = fst oo inst;
|
haftmann@25462
|
179 |
|
wenzelm@25020
|
180 |
fun inst_thms thy = (Symtab.fold (Symtab.fold (cons o snd o snd) o snd) o fst)
|
haftmann@25462
|
181 |
(InstData.get thy) [];
|
haftmann@25462
|
182 |
|
haftmann@25462
|
183 |
val param_const = Symtab.lookup o snd o InstData.get;
|
haftmann@25462
|
184 |
|
haftmann@24589
|
185 |
fun add_inst (c, tyco) inst = (InstData.map o apfst
|
haftmann@24589
|
186 |
o Symtab.map_default (c, Symtab.empty)) (Symtab.update_new (tyco, inst))
|
haftmann@24423
|
187 |
#> (InstData.map o apsnd) (Symtab.update_new (fst inst, (c, tyco)));
|
haftmann@24304
|
188 |
|
wenzelm@25020
|
189 |
fun unoverload thy = MetaSimplifier.rewrite true (inst_thms thy);
|
wenzelm@25020
|
190 |
fun overload thy = MetaSimplifier.rewrite true (map Thm.symmetric (inst_thms thy));
|
haftmann@24304
|
191 |
|
haftmann@24423
|
192 |
fun unoverload_const thy (c_ty as (c, _)) =
|
haftmann@24423
|
193 |
case AxClass.class_of_param thy c
|
haftmann@25462
|
194 |
of SOME class => (case inst_tyco thy c_ty
|
haftmann@25462
|
195 |
of SOME tyco => (case try (inst thy) (c, tyco)
|
haftmann@24423
|
196 |
of SOME (c, _) => c
|
haftmann@24423
|
197 |
| NONE => c)
|
haftmann@25462
|
198 |
| NONE => c)
|
haftmann@24423
|
199 |
| NONE => c;
|
haftmann@24423
|
200 |
|
haftmann@25462
|
201 |
|
haftmann@25462
|
202 |
(* declaration and definition of instances of overloaded constants *)
|
haftmann@25462
|
203 |
|
haftmann@25462
|
204 |
fun primitive_note kind (name, thm) =
|
haftmann@25462
|
205 |
PureThy.note_thmss_i kind [((name, []), [([thm], [])])]
|
haftmann@25462
|
206 |
#>> (fn [(_, [thm])] => thm);
|
haftmann@25462
|
207 |
|
haftmann@25462
|
208 |
fun declare_overloaded (c, ty, mx) thy =
|
haftmann@25462
|
209 |
let
|
haftmann@25462
|
210 |
val SOME class = AxClass.class_of_param thy c;
|
haftmann@25462
|
211 |
val SOME tyco = inst_tyco thy (c, ty);
|
haftmann@25462
|
212 |
val name_inst = NameSpace.base class ^ "_" ^ NameSpace.base tyco ^ "_inst";
|
haftmann@25462
|
213 |
val c' = NameSpace.base c;
|
haftmann@25462
|
214 |
val ty' = Type.strip_sorts ty;
|
haftmann@25462
|
215 |
in
|
haftmann@25462
|
216 |
thy
|
haftmann@25462
|
217 |
|> Sign.sticky_prefix name_inst
|
haftmann@25462
|
218 |
|> Sign.no_base_names
|
haftmann@25462
|
219 |
|> Sign.notation true Syntax.mode_default [(Const (c, ty), mx)]
|
haftmann@25462
|
220 |
|> Sign.declare_const [] (c', ty', NoSyn)
|
haftmann@25462
|
221 |
|-> (fn const' as Const (c'', _) => Thm.add_def true
|
haftmann@25462
|
222 |
(Thm.def_name c', Logic.mk_equals (Const (c, ty'), const'))
|
haftmann@25462
|
223 |
#>> Thm.varifyT
|
haftmann@25462
|
224 |
#-> (fn thm => add_inst (c, tyco) (c'', thm)
|
haftmann@25462
|
225 |
#> primitive_note Thm.internalK (c', thm)
|
haftmann@25462
|
226 |
#> snd
|
haftmann@25462
|
227 |
#> Sign.restore_naming thy
|
haftmann@25462
|
228 |
#> pair (Const (c, ty))))
|
haftmann@25462
|
229 |
end;
|
haftmann@25462
|
230 |
|
haftmann@25462
|
231 |
fun define_overloaded name (c, t) thy =
|
haftmann@25462
|
232 |
let
|
haftmann@25462
|
233 |
val ty = Term.fastype_of t;
|
haftmann@25462
|
234 |
val SOME tyco = inst_tyco thy (c, ty);
|
haftmann@25462
|
235 |
val (c', eq) = inst thy (c, tyco);
|
haftmann@25462
|
236 |
val [Type (_, tys)] = Sign.const_typargs thy (c, ty);
|
haftmann@25462
|
237 |
val eq' = eq
|
haftmann@25462
|
238 |
|> Drule.instantiate' (map (SOME o Thm.ctyp_of thy) tys) [];
|
haftmann@25462
|
239 |
(*FIXME proper recover_sort mechanism*)
|
haftmann@25462
|
240 |
val prop = Logic.mk_equals (Const (c', ty), t);
|
haftmann@25462
|
241 |
val name' = if name = "" then
|
haftmann@25462
|
242 |
Thm.def_name (NameSpace.base c ^ "_" ^ NameSpace.base tyco) else name;
|
haftmann@25462
|
243 |
in
|
haftmann@25462
|
244 |
thy
|
haftmann@25462
|
245 |
|> Thm.add_def false (name', prop)
|
haftmann@25462
|
246 |
|>> (fn thm => Thm.transitive eq' thm)
|
haftmann@25462
|
247 |
end;
|
haftmann@25462
|
248 |
|
haftmann@25462
|
249 |
|
haftmann@25462
|
250 |
(* legacy *)
|
haftmann@24304
|
251 |
|
haftmann@24304
|
252 |
fun add_inst_def (class, tyco) (c, ty) thy =
|
haftmann@24304
|
253 |
let
|
wenzelm@25024
|
254 |
val tyco_base = Sign.base_name tyco;
|
wenzelm@25024
|
255 |
val name_inst = Sign.base_name class ^ "_" ^ tyco_base ^ "_inst";
|
wenzelm@25024
|
256 |
val c_inst_base = Sign.base_name c ^ "_" ^ tyco_base;
|
haftmann@24304
|
257 |
in
|
haftmann@24304
|
258 |
thy
|
haftmann@24304
|
259 |
|> Sign.sticky_prefix name_inst
|
wenzelm@25020
|
260 |
|> Sign.declare_const [] (c_inst_base, ty, NoSyn)
|
wenzelm@24968
|
261 |
|-> (fn const as Const (c_inst, _) =>
|
wenzelm@24968
|
262 |
PureThy.add_defs_i false
|
wenzelm@24968
|
263 |
[((Thm.def_name c_inst_base, Logic.mk_equals (const, Const (c, ty))), [])]
|
wenzelm@24968
|
264 |
#-> (fn [thm] => add_inst (c, tyco) (c_inst, Thm.symmetric thm)))
|
haftmann@24435
|
265 |
|> Sign.restore_naming thy
|
haftmann@24304
|
266 |
end;
|
haftmann@24304
|
267 |
|
haftmann@24304
|
268 |
fun add_inst_def' (class, tyco) (c, ty) thy =
|
haftmann@24423
|
269 |
if case Symtab.lookup (fst (InstData.get thy)) c
|
haftmann@24304
|
270 |
of NONE => true
|
haftmann@24304
|
271 |
| SOME tab => is_none (Symtab.lookup tab tyco)
|
haftmann@24304
|
272 |
then add_inst_def (class, tyco) (c, Logic.unvarifyT ty) thy
|
haftmann@24304
|
273 |
else thy;
|
haftmann@24304
|
274 |
|
haftmann@24304
|
275 |
fun add_def ((class, tyco), ((name, prop), atts)) thy =
|
haftmann@24304
|
276 |
let
|
haftmann@24589
|
277 |
val ((lhs as Const (c, ty), args), rhs) =
|
haftmann@24589
|
278 |
(apfst Term.strip_comb o Logic.dest_equals) prop;
|
haftmann@24304
|
279 |
in
|
haftmann@24304
|
280 |
thy
|
wenzelm@25020
|
281 |
|> PureThy.add_defs_i true [((name, prop), map (Attrib.attribute_i thy) atts)]
|
haftmann@25462
|
282 |
|-> (fn [def] => add_inst_def (class, tyco) (c, ty) #> pair def)
|
haftmann@24304
|
283 |
end;
|
haftmann@24304
|
284 |
|
haftmann@24304
|
285 |
|
haftmann@24589
|
286 |
(** instances with implicit parameter handling **)
|
haftmann@24218
|
287 |
|
haftmann@24218
|
288 |
local
|
haftmann@24218
|
289 |
|
wenzelm@24707
|
290 |
fun gen_read_def thy prep_att parse_prop ((raw_name, raw_atts), raw_t) =
|
haftmann@24218
|
291 |
let
|
wenzelm@24707
|
292 |
val ctxt = ProofContext.init thy;
|
wenzelm@24707
|
293 |
val t = parse_prop ctxt raw_t |> Syntax.check_prop ctxt;
|
wenzelm@24981
|
294 |
val ((c, ty), _) = Sign.cert_def ctxt t;
|
haftmann@24218
|
295 |
val atts = map (prep_att thy) raw_atts;
|
haftmann@24218
|
296 |
val insts = Consts.typargs (Sign.consts_of thy) (c, ty);
|
haftmann@24218
|
297 |
val name = case raw_name
|
haftmann@24218
|
298 |
of "" => NONE
|
haftmann@24218
|
299 |
| _ => SOME raw_name;
|
haftmann@24218
|
300 |
in (c, (insts, ((name, t), atts))) end;
|
haftmann@24218
|
301 |
|
wenzelm@24707
|
302 |
fun read_def_cmd thy = gen_read_def thy Attrib.intern_src Syntax.parse_prop;
|
haftmann@24218
|
303 |
fun read_def thy = gen_read_def thy (K I) (K I);
|
haftmann@24218
|
304 |
|
haftmann@24589
|
305 |
fun gen_instance prep_arity read_def do_proof raw_arities raw_defs after_qed theory =
|
haftmann@24218
|
306 |
let
|
haftmann@24218
|
307 |
val arities = map (prep_arity theory) raw_arities;
|
wenzelm@25020
|
308 |
val _ = if null arities then error "At least one arity must be given" else ();
|
haftmann@24218
|
309 |
val _ = case (duplicates (op =) o map #1) arities
|
haftmann@24218
|
310 |
of [] => ()
|
wenzelm@25020
|
311 |
| dupl_tycos => error ("Type constructors occur more than once in arities: "
|
wenzelm@25020
|
312 |
^ commas_quote dupl_tycos);
|
haftmann@24218
|
313 |
fun get_consts_class tyco ty class =
|
haftmann@24218
|
314 |
let
|
wenzelm@24968
|
315 |
val cs = (these o try (#params o AxClass.get_info theory)) class;
|
haftmann@24218
|
316 |
val subst_ty = map_type_tfree (K ty);
|
haftmann@24218
|
317 |
in
|
haftmann@24304
|
318 |
map (fn (c, ty) => (c, ((class, tyco), subst_ty ty))) cs
|
haftmann@24218
|
319 |
end;
|
haftmann@24218
|
320 |
fun get_consts_sort (tyco, asorts, sort) =
|
haftmann@24218
|
321 |
let
|
haftmann@24589
|
322 |
val ty = Type (tyco, map (fn (v, sort) => TVar ((v, 0), sort))
|
wenzelm@24847
|
323 |
(Name.names Name.context Name.aT asorts))
|
wenzelm@24731
|
324 |
in maps (get_consts_class tyco ty) (Sign.complete_sort theory sort) end;
|
haftmann@24218
|
325 |
val cs = maps get_consts_sort arities;
|
haftmann@24218
|
326 |
fun mk_typnorm thy (ty, ty_sc) =
|
haftmann@24218
|
327 |
case try (Sign.typ_match thy (Logic.varifyT ty_sc, ty)) Vartab.empty
|
haftmann@24218
|
328 |
of SOME env => SOME (Logic.varifyT #> Envir.typ_subst_TVars env #> Logic.unvarifyT)
|
haftmann@24218
|
329 |
| NONE => NONE;
|
haftmann@24218
|
330 |
fun read_defs defs cs thy_read =
|
haftmann@24218
|
331 |
let
|
haftmann@24218
|
332 |
fun read raw_def cs =
|
haftmann@24218
|
333 |
let
|
haftmann@24218
|
334 |
val (c, (inst, ((name_opt, t), atts))) = read_def thy_read raw_def;
|
haftmann@24218
|
335 |
val ty = Consts.instance (Sign.consts_of thy_read) (c, inst);
|
haftmann@24304
|
336 |
val ((class, tyco), ty') = case AList.lookup (op =) cs c
|
wenzelm@25020
|
337 |
of NONE => error ("Illegal definition for constant " ^ quote c)
|
haftmann@24218
|
338 |
| SOME class_ty => class_ty;
|
haftmann@24218
|
339 |
val name = case name_opt
|
haftmann@24218
|
340 |
of NONE => Thm.def_name (Logic.name_arity (tyco, [], c))
|
haftmann@24218
|
341 |
| SOME name => name;
|
haftmann@24218
|
342 |
val t' = case mk_typnorm thy_read (ty', ty)
|
wenzelm@25020
|
343 |
of NONE => error ("Illegal definition for constant " ^
|
haftmann@24218
|
344 |
quote (c ^ "::" ^ setmp show_sorts true
|
haftmann@24218
|
345 |
(Sign.string_of_typ thy_read) ty))
|
haftmann@24218
|
346 |
| SOME norm => map_types norm t
|
haftmann@24218
|
347 |
in (((class, tyco), ((name, t'), atts)), AList.delete (op =) c cs) end;
|
haftmann@24218
|
348 |
in fold_map read defs cs end;
|
haftmann@24304
|
349 |
val (defs, other_cs) = read_defs raw_defs cs
|
haftmann@24218
|
350 |
(fold Sign.primitive_arity arities (Theory.copy theory));
|
haftmann@24423
|
351 |
fun after_qed' cs defs =
|
wenzelm@24766
|
352 |
fold Sign.add_const_constraint (map (apsnd SOME) cs)
|
haftmann@24423
|
353 |
#> after_qed defs;
|
haftmann@24218
|
354 |
in
|
haftmann@24218
|
355 |
theory
|
haftmann@25038
|
356 |
|> fold_map get_remove_global_constraint (map fst cs |> distinct (op =))
|
haftmann@24304
|
357 |
||>> fold_map add_def defs
|
haftmann@24304
|
358 |
||> fold (fn (c, ((class, tyco), ty)) => add_inst_def' (class, tyco) (c, ty)) other_cs
|
haftmann@24423
|
359 |
|-> (fn (cs, defs) => do_proof (after_qed' cs defs) arities defs)
|
haftmann@24218
|
360 |
end;
|
haftmann@24218
|
361 |
|
haftmann@24423
|
362 |
fun tactic_proof tac f arities defs =
|
haftmann@24218
|
363 |
fold (fn arity => AxClass.prove_arity arity tac) arities
|
haftmann@24423
|
364 |
#> f
|
haftmann@24423
|
365 |
#> pair defs;
|
haftmann@24218
|
366 |
|
haftmann@24218
|
367 |
in
|
haftmann@24218
|
368 |
|
haftmann@24589
|
369 |
val instance = gen_instance Sign.cert_arity read_def
|
haftmann@24589
|
370 |
(fn f => fn arities => fn defs => instance_arity f arities);
|
haftmann@24589
|
371 |
val instance_cmd = gen_instance Sign.read_arity read_def_cmd
|
haftmann@24589
|
372 |
(fn f => fn arities => fn defs => instance_arity f arities);
|
haftmann@24589
|
373 |
fun prove_instance tac arities defs =
|
haftmann@24589
|
374 |
gen_instance Sign.cert_arity read_def
|
haftmann@24589
|
375 |
(tactic_proof tac) arities defs (K I);
|
haftmann@24218
|
376 |
|
haftmann@24218
|
377 |
end; (*local*)
|
haftmann@24218
|
378 |
|
haftmann@24218
|
379 |
|
haftmann@24218
|
380 |
|
haftmann@24589
|
381 |
(** class data **)
|
haftmann@24218
|
382 |
|
haftmann@24218
|
383 |
datatype class_data = ClassData of {
|
haftmann@24218
|
384 |
consts: (string * string) list
|
haftmann@24836
|
385 |
(*locale parameter ~> constant name*),
|
haftmann@25062
|
386 |
base_sort: sort,
|
haftmann@25083
|
387 |
inst: term option list
|
haftmann@25083
|
388 |
(*canonical interpretation*),
|
haftmann@25062
|
389 |
morphism: morphism,
|
haftmann@25062
|
390 |
(*partial morphism of canonical interpretation*)
|
haftmann@24657
|
391 |
intro: thm,
|
haftmann@24657
|
392 |
defs: thm list,
|
haftmann@25368
|
393 |
operations: (string * (class * (typ * term))) list
|
haftmann@24657
|
394 |
};
|
haftmann@24218
|
395 |
|
haftmann@24657
|
396 |
fun rep_class_data (ClassData d) = d;
|
haftmann@25062
|
397 |
fun mk_class_data ((consts, base_sort, inst, morphism, intro),
|
haftmann@25368
|
398 |
(defs, operations)) =
|
haftmann@25062
|
399 |
ClassData { consts = consts, base_sort = base_sort, inst = inst,
|
haftmann@25062
|
400 |
morphism = morphism, intro = intro, defs = defs,
|
haftmann@25368
|
401 |
operations = operations };
|
haftmann@25062
|
402 |
fun map_class_data f (ClassData { consts, base_sort, inst, morphism, intro,
|
haftmann@25368
|
403 |
defs, operations }) =
|
haftmann@25062
|
404 |
mk_class_data (f ((consts, base_sort, inst, morphism, intro),
|
haftmann@25368
|
405 |
(defs, operations)));
|
haftmann@25038
|
406 |
fun merge_class_data _ (ClassData { consts = consts,
|
haftmann@25062
|
407 |
base_sort = base_sort, inst = inst, morphism = morphism, intro = intro,
|
haftmann@25368
|
408 |
defs = defs1, operations = operations1 },
|
haftmann@25062
|
409 |
ClassData { consts = _, base_sort = _, inst = _, morphism = _, intro = _,
|
haftmann@25368
|
410 |
defs = defs2, operations = operations2 }) =
|
haftmann@25062
|
411 |
mk_class_data ((consts, base_sort, inst, morphism, intro),
|
haftmann@24914
|
412 |
(Thm.merge_thms (defs1, defs2),
|
haftmann@25368
|
413 |
AList.merge (op =) (K true) (operations1, operations2)));
|
haftmann@24218
|
414 |
|
haftmann@24218
|
415 |
structure ClassData = TheoryDataFun
|
haftmann@24218
|
416 |
(
|
haftmann@25038
|
417 |
type T = class_data Graph.T
|
haftmann@25038
|
418 |
val empty = Graph.empty;
|
haftmann@24218
|
419 |
val copy = I;
|
haftmann@24218
|
420 |
val extend = I;
|
haftmann@25038
|
421 |
fun merge _ = Graph.join merge_class_data;
|
haftmann@24218
|
422 |
);
|
haftmann@24218
|
423 |
|
haftmann@24218
|
424 |
|
haftmann@24218
|
425 |
(* queries *)
|
haftmann@24218
|
426 |
|
haftmann@25038
|
427 |
val lookup_class_data = Option.map rep_class_data oo try o Graph.get_node o ClassData.get;
|
haftmann@24218
|
428 |
|
haftmann@24589
|
429 |
fun the_class_data thy class = case lookup_class_data thy class
|
wenzelm@25020
|
430 |
of NONE => error ("Undeclared class " ^ quote class)
|
haftmann@24589
|
431 |
| SOME data => data;
|
haftmann@24218
|
432 |
|
haftmann@25038
|
433 |
val is_class = is_some oo lookup_class_data;
|
haftmann@25038
|
434 |
|
haftmann@25038
|
435 |
val ancestry = Graph.all_succs o ClassData.get;
|
haftmann@24218
|
436 |
|
haftmann@25002
|
437 |
fun these_params thy =
|
haftmann@24218
|
438 |
let
|
haftmann@24218
|
439 |
fun params class =
|
haftmann@24218
|
440 |
let
|
wenzelm@24930
|
441 |
val const_typs = (#params o AxClass.get_info thy) class;
|
haftmann@24657
|
442 |
val const_names = (#consts o the_class_data thy) class;
|
haftmann@24218
|
443 |
in
|
haftmann@24218
|
444 |
(map o apsnd) (fn c => (c, (the o AList.lookup (op =) const_typs) c)) const_names
|
haftmann@24218
|
445 |
end;
|
haftmann@24218
|
446 |
in maps params o ancestry thy end;
|
haftmann@24218
|
447 |
|
haftmann@24657
|
448 |
fun these_defs thy = maps (these o Option.map #defs o lookup_class_data thy) o ancestry thy;
|
haftmann@24218
|
449 |
|
haftmann@25062
|
450 |
fun morphism thy = #morphism o the_class_data thy;
|
haftmann@25062
|
451 |
|
haftmann@24218
|
452 |
fun these_intros thy =
|
haftmann@24657
|
453 |
Graph.fold (fn (_, (data, _)) => insert Thm.eq_thm ((#intro o rep_class_data) data))
|
haftmann@25038
|
454 |
(ClassData.get thy) [];
|
haftmann@24218
|
455 |
|
haftmann@24836
|
456 |
fun these_operations thy =
|
haftmann@24836
|
457 |
maps (#operations o the_class_data thy) o ancestry thy;
|
haftmann@24657
|
458 |
|
haftmann@24218
|
459 |
fun print_classes thy =
|
haftmann@24218
|
460 |
let
|
wenzelm@24920
|
461 |
val ctxt = ProofContext.init thy;
|
haftmann@24218
|
462 |
val algebra = Sign.classes_of thy;
|
haftmann@24218
|
463 |
val arities =
|
haftmann@24218
|
464 |
Symtab.empty
|
haftmann@24218
|
465 |
|> Symtab.fold (fn (tyco, arities) => fold (fn (class, _) =>
|
haftmann@24218
|
466 |
Symtab.map_default (class, []) (insert (op =) tyco)) arities)
|
haftmann@24218
|
467 |
((#arities o Sorts.rep_algebra) algebra);
|
haftmann@24218
|
468 |
val the_arities = these o Symtab.lookup arities;
|
haftmann@24218
|
469 |
fun mk_arity class tyco =
|
haftmann@24218
|
470 |
let
|
haftmann@24218
|
471 |
val Ss = Sorts.mg_domain algebra tyco [class];
|
wenzelm@24920
|
472 |
in Syntax.pretty_arity ctxt (tyco, Ss, [class]) end;
|
haftmann@24218
|
473 |
fun mk_param (c, ty) = Pretty.str (Sign.extern_const thy c ^ " :: "
|
wenzelm@24920
|
474 |
^ setmp show_sorts false (Syntax.string_of_typ ctxt o Type.strip_sorts) ty);
|
haftmann@24218
|
475 |
fun mk_entry class = (Pretty.block o Pretty.fbreaks o map_filter I) [
|
haftmann@25062
|
476 |
(SOME o Pretty.str) ("class " ^ Sign.extern_class thy class ^ ":"),
|
haftmann@24218
|
477 |
(SOME o Pretty.block) [Pretty.str "supersort: ",
|
wenzelm@24920
|
478 |
(Syntax.pretty_sort ctxt o Sign.minimize_sort thy o Sign.super_classes thy) class],
|
haftmann@25062
|
479 |
if is_class thy class then (SOME o Pretty.str)
|
haftmann@25062
|
480 |
("locale: " ^ Locale.extern thy class) else NONE,
|
haftmann@25062
|
481 |
((fn [] => NONE | ps => (SOME o Pretty.block o Pretty.fbreaks)
|
haftmann@25062
|
482 |
(Pretty.str "parameters:" :: ps)) o map mk_param
|
wenzelm@24930
|
483 |
o these o Option.map #params o try (AxClass.get_info thy)) class,
|
haftmann@24218
|
484 |
(SOME o Pretty.block o Pretty.breaks) [
|
haftmann@24218
|
485 |
Pretty.str "instances:",
|
haftmann@24218
|
486 |
Pretty.list "" "" (map (mk_arity class) (the_arities class))
|
haftmann@24218
|
487 |
]
|
haftmann@24218
|
488 |
]
|
haftmann@24218
|
489 |
in
|
haftmann@24589
|
490 |
(Pretty.writeln o Pretty.chunks o separate (Pretty.str "")
|
haftmann@24589
|
491 |
o map mk_entry o Sorts.all_classes) algebra
|
haftmann@24218
|
492 |
end;
|
haftmann@24218
|
493 |
|
haftmann@24218
|
494 |
|
haftmann@24218
|
495 |
(* updaters *)
|
haftmann@24218
|
496 |
|
haftmann@25163
|
497 |
fun add_class_data ((class, superclasses), (cs, base_sort, inst, phi, intro)) thy =
|
haftmann@25002
|
498 |
let
|
haftmann@25368
|
499 |
val operations = map (fn (v_ty as (_, ty), (c, _)) =>
|
haftmann@25368
|
500 |
(c, (class, (ty, Free v_ty)))) cs;
|
haftmann@25002
|
501 |
val cs = (map o pairself) fst cs;
|
haftmann@25038
|
502 |
val add_class = Graph.new_node (class,
|
haftmann@25368
|
503 |
mk_class_data ((cs, base_sort, map (SOME o Const) inst, phi, intro), ([], operations)))
|
haftmann@25002
|
504 |
#> fold (curry Graph.add_edge class) superclasses;
|
haftmann@25002
|
505 |
in
|
haftmann@25038
|
506 |
ClassData.map add_class thy
|
haftmann@25002
|
507 |
end;
|
haftmann@24218
|
508 |
|
haftmann@25368
|
509 |
fun register_operation class (c, (t, some_def)) thy =
|
haftmann@25062
|
510 |
let
|
haftmann@25368
|
511 |
val base_sort = (#base_sort o the_class_data thy) class;
|
haftmann@25239
|
512 |
val prep_typ = map_atyps
|
haftmann@25368
|
513 |
(fn TVar (vi as (v, _), sort) => if Name.aT = v
|
haftmann@25368
|
514 |
then TFree (v, base_sort) else TVar (vi, sort));
|
haftmann@25368
|
515 |
val t' = map_types prep_typ t;
|
haftmann@25368
|
516 |
val ty' = Term.fastype_of t';
|
haftmann@25062
|
517 |
in
|
haftmann@25062
|
518 |
thy
|
haftmann@25062
|
519 |
|> (ClassData.map o Graph.map_node class o map_class_data o apsnd)
|
haftmann@25368
|
520 |
(fn (defs, operations) =>
|
haftmann@25096
|
521 |
(fold cons (the_list some_def) defs,
|
haftmann@25368
|
522 |
(c, (class, (ty', t'))) :: operations))
|
haftmann@25062
|
523 |
end;
|
haftmann@24218
|
524 |
|
haftmann@24589
|
525 |
|
haftmann@24589
|
526 |
(** rule calculation, tactics and methods **)
|
haftmann@24589
|
527 |
|
wenzelm@25024
|
528 |
val class_prefix = Logic.const_of_class o Sign.base_name;
|
wenzelm@25024
|
529 |
|
haftmann@25062
|
530 |
fun calculate_morphism class cs =
|
haftmann@25062
|
531 |
let
|
haftmann@25062
|
532 |
val subst_typ = Term.map_type_tfree (fn var as (v, sort) =>
|
haftmann@25062
|
533 |
if v = Name.aT then TVar ((v, 0), [class]) else TVar ((v, 0), sort));
|
haftmann@25062
|
534 |
fun subst_aterm (t as Free (v, ty)) = (case AList.lookup (op =) cs v
|
haftmann@25062
|
535 |
of SOME (c, _) => Const (c, ty)
|
haftmann@25062
|
536 |
| NONE => t)
|
haftmann@25062
|
537 |
| subst_aterm t = t;
|
haftmann@25062
|
538 |
val subst_term = map_aterms subst_aterm #> map_types subst_typ;
|
haftmann@25062
|
539 |
in
|
haftmann@25209
|
540 |
Morphism.term_morphism subst_term
|
haftmann@25062
|
541 |
$> Morphism.typ_morphism subst_typ
|
haftmann@25062
|
542 |
end;
|
haftmann@25062
|
543 |
|
haftmann@25038
|
544 |
fun class_intro thy class sups =
|
haftmann@24589
|
545 |
let
|
haftmann@24589
|
546 |
fun class_elim class =
|
wenzelm@25020
|
547 |
case (#axioms o AxClass.get_info thy) class
|
wenzelm@25020
|
548 |
of [thm] => SOME (Drule.unconstrainTs thm)
|
haftmann@24589
|
549 |
| [] => NONE;
|
haftmann@25038
|
550 |
val pred_intro = case Locale.intros thy class
|
haftmann@24589
|
551 |
of ([ax_intro], [intro]) => intro |> OF_LAST ax_intro |> SOME
|
haftmann@24589
|
552 |
| ([intro], []) => SOME intro
|
haftmann@24589
|
553 |
| ([], [intro]) => SOME intro
|
haftmann@24589
|
554 |
| _ => NONE;
|
haftmann@24589
|
555 |
val pred_intro' = pred_intro
|
haftmann@24589
|
556 |
|> Option.map (fn intro => intro OF map_filter class_elim sups);
|
wenzelm@24930
|
557 |
val class_intro = (#intro o AxClass.get_info thy) class;
|
haftmann@24589
|
558 |
val raw_intro = case pred_intro'
|
haftmann@24589
|
559 |
of SOME pred_intro => class_intro |> OF_LAST pred_intro
|
haftmann@24589
|
560 |
| NONE => class_intro;
|
haftmann@24589
|
561 |
val sort = Sign.super_classes thy class;
|
wenzelm@24847
|
562 |
val typ = TVar ((Name.aT, 0), sort);
|
haftmann@24589
|
563 |
val defs = these_defs thy sups;
|
haftmann@24589
|
564 |
in
|
haftmann@24589
|
565 |
raw_intro
|
haftmann@24589
|
566 |
|> Drule.instantiate' [SOME (Thm.ctyp_of thy typ)] []
|
haftmann@24589
|
567 |
|> strip_all_ofclass thy sort
|
haftmann@24589
|
568 |
|> Thm.strip_shyps
|
haftmann@24589
|
569 |
|> MetaSimplifier.rewrite_rule defs
|
haftmann@24589
|
570 |
|> Drule.unconstrainTs
|
haftmann@24589
|
571 |
end;
|
haftmann@24589
|
572 |
|
haftmann@24589
|
573 |
fun class_interpretation class facts defs thy =
|
haftmann@24589
|
574 |
let
|
haftmann@25038
|
575 |
val params = these_params thy [class];
|
haftmann@25083
|
576 |
val inst = (#inst o the_class_data thy) class;
|
wenzelm@25020
|
577 |
val tac = ALLGOALS (ProofContext.fact_tac facts);
|
haftmann@25038
|
578 |
val prfx = class_prefix class;
|
haftmann@24589
|
579 |
in
|
haftmann@25038
|
580 |
thy
|
haftmann@25038
|
581 |
|> fold_map (get_remove_global_constraint o fst o snd) params
|
ballarin@25094
|
582 |
||> prove_interpretation tac ((false, prfx), []) (Locale.Locale class)
|
ballarin@25094
|
583 |
(inst, map (fn def => (("", []), def)) defs)
|
haftmann@25038
|
584 |
|-> (fn cs => fold (Sign.add_const_constraint o apsnd SOME) cs)
|
haftmann@24589
|
585 |
end;
|
haftmann@24218
|
586 |
|
haftmann@24218
|
587 |
fun intro_classes_tac facts st =
|
haftmann@24218
|
588 |
let
|
haftmann@24218
|
589 |
val thy = Thm.theory_of_thm st;
|
haftmann@24218
|
590 |
val classes = Sign.all_classes thy;
|
haftmann@24218
|
591 |
val class_trivs = map (Thm.class_triv thy) classes;
|
haftmann@24218
|
592 |
val class_intros = these_intros thy;
|
wenzelm@24930
|
593 |
val axclass_intros = map_filter (try (#intro o AxClass.get_info thy)) classes;
|
haftmann@24218
|
594 |
in
|
haftmann@25268
|
595 |
Method.intros_tac (class_trivs @ class_intros @ axclass_intros) facts st
|
haftmann@24218
|
596 |
end;
|
haftmann@24218
|
597 |
|
haftmann@24218
|
598 |
fun default_intro_classes_tac [] = intro_classes_tac []
|
wenzelm@24930
|
599 |
| default_intro_classes_tac _ = no_tac;
|
haftmann@24218
|
600 |
|
haftmann@24218
|
601 |
fun default_tac rules ctxt facts =
|
haftmann@24218
|
602 |
HEADGOAL (Method.some_rule_tac rules ctxt facts) ORELSE
|
haftmann@24218
|
603 |
default_intro_classes_tac facts;
|
haftmann@24218
|
604 |
|
haftmann@24218
|
605 |
val _ = Context.add_setup (Method.add_methods
|
haftmann@24218
|
606 |
[("intro_classes", Method.no_args (Method.METHOD intro_classes_tac),
|
haftmann@24218
|
607 |
"back-chain introduction rules of classes"),
|
haftmann@24218
|
608 |
("default", Method.thms_ctxt_args (Method.METHOD oo default_tac),
|
haftmann@24218
|
609 |
"apply some intro/elim rule")]);
|
haftmann@24218
|
610 |
|
haftmann@25195
|
611 |
fun subclass_rule thy (sub, sup) =
|
haftmann@25195
|
612 |
let
|
haftmann@25195
|
613 |
val ctxt = Locale.init sub thy;
|
haftmann@25195
|
614 |
val ctxt_thy = ProofContext.init thy;
|
haftmann@25195
|
615 |
val props =
|
haftmann@25195
|
616 |
Locale.global_asms_of thy sup
|
haftmann@25195
|
617 |
|> maps snd
|
haftmann@25195
|
618 |
|> map (ObjectLogic.ensure_propT thy);
|
haftmann@25195
|
619 |
fun tac { prems, context } =
|
haftmann@25195
|
620 |
Locale.intro_locales_tac true context prems
|
haftmann@25195
|
621 |
ORELSE ALLGOALS assume_tac;
|
haftmann@25195
|
622 |
in
|
haftmann@25195
|
623 |
Goal.prove_multi ctxt [] [] props tac
|
haftmann@25195
|
624 |
|> map (Assumption.export false ctxt ctxt_thy)
|
haftmann@25195
|
625 |
|> Variable.export ctxt ctxt_thy
|
haftmann@25195
|
626 |
end;
|
haftmann@25195
|
627 |
|
haftmann@25195
|
628 |
fun prove_single_subclass (sub, sup) thms ctxt thy =
|
haftmann@25195
|
629 |
let
|
haftmann@25195
|
630 |
val ctxt_thy = ProofContext.init thy;
|
haftmann@25195
|
631 |
val subclass_rule = Conjunction.intr_balanced thms
|
haftmann@25195
|
632 |
|> Assumption.export false ctxt ctxt_thy
|
haftmann@25195
|
633 |
|> singleton (Variable.export ctxt ctxt_thy);
|
haftmann@25195
|
634 |
val sub_inst = Thm.ctyp_of thy (TVar ((Name.aT, 0), [sub]));
|
haftmann@25195
|
635 |
val sub_ax = #axioms (AxClass.get_info thy sub);
|
haftmann@25195
|
636 |
val classrel =
|
haftmann@25195
|
637 |
#intro (AxClass.get_info thy sup)
|
haftmann@25195
|
638 |
|> Drule.instantiate' [SOME sub_inst] []
|
haftmann@25195
|
639 |
|> OF_LAST (subclass_rule OF sub_ax)
|
haftmann@25195
|
640 |
|> strip_all_ofclass thy (Sign.super_classes thy sup)
|
haftmann@25195
|
641 |
|> Thm.strip_shyps
|
haftmann@25195
|
642 |
in
|
haftmann@25195
|
643 |
thy
|
haftmann@25195
|
644 |
|> AxClass.add_classrel classrel
|
haftmann@25195
|
645 |
|> prove_interpretation_in (ALLGOALS (ProofContext.fact_tac thms))
|
haftmann@25195
|
646 |
I (sub, Locale.Locale sup)
|
haftmann@25195
|
647 |
|> ClassData.map (Graph.add_edge (sub, sup))
|
haftmann@25195
|
648 |
end;
|
haftmann@25195
|
649 |
|
haftmann@25195
|
650 |
fun prove_subclass (sub, sup) thms ctxt thy =
|
haftmann@25195
|
651 |
let
|
haftmann@25268
|
652 |
val classes = ClassData.get thy;
|
haftmann@25268
|
653 |
val is_sup = not o null o curry (Graph.irreducible_paths classes) sub;
|
haftmann@25268
|
654 |
val supclasses = Graph.all_succs classes [sup] |> filter_out is_sup;
|
haftmann@25195
|
655 |
fun transform sup' = subclass_rule thy (sup, sup') |> map (fn thm => thm OF thms);
|
haftmann@25195
|
656 |
in
|
haftmann@25195
|
657 |
thy
|
haftmann@25195
|
658 |
|> fold_rev (fn sup' => prove_single_subclass (sub, sup')
|
haftmann@25195
|
659 |
(transform sup') ctxt) supclasses
|
haftmann@25195
|
660 |
end;
|
haftmann@25195
|
661 |
|
haftmann@24218
|
662 |
|
haftmann@24589
|
663 |
(** classes and class target **)
|
haftmann@24218
|
664 |
|
haftmann@25002
|
665 |
(* class context syntax *)
|
haftmann@24748
|
666 |
|
haftmann@25083
|
667 |
structure ClassSyntax = ProofDataFun(
|
haftmann@25083
|
668 |
type T = {
|
haftmann@25368
|
669 |
local_constraints: (string * typ) list,
|
haftmann@25368
|
670 |
global_constraints: (string * typ) list,
|
haftmann@25083
|
671 |
base_sort: sort,
|
haftmann@25368
|
672 |
operations: (string * (typ * term)) list,
|
haftmann@25195
|
673 |
unchecks: (term * term) list,
|
haftmann@25083
|
674 |
passed: bool
|
haftmann@25368
|
675 |
};
|
haftmann@25368
|
676 |
fun init _ = {
|
haftmann@25368
|
677 |
local_constraints = [],
|
haftmann@25368
|
678 |
global_constraints = [],
|
haftmann@25368
|
679 |
base_sort = [],
|
haftmann@25368
|
680 |
operations = [],
|
haftmann@25368
|
681 |
unchecks = [],
|
haftmann@25368
|
682 |
passed = true
|
haftmann@25368
|
683 |
};;
|
haftmann@25083
|
684 |
);
|
haftmann@25083
|
685 |
|
wenzelm@25344
|
686 |
fun synchronize_syntax sups base_sort ctxt =
|
haftmann@24914
|
687 |
let
|
wenzelm@25344
|
688 |
val thy = ProofContext.theory_of ctxt;
|
haftmann@25368
|
689 |
fun subst_class_typ sort = map_atyps
|
haftmann@25368
|
690 |
(fn TFree _ => TVar ((Name.aT, 0), sort) | ty' => ty');
|
haftmann@25083
|
691 |
val operations = these_operations thy sups;
|
haftmann@25368
|
692 |
val local_constraints =
|
haftmann@25368
|
693 |
(map o apsnd) (subst_class_typ base_sort o fst o snd) operations;
|
haftmann@25368
|
694 |
val global_constraints =
|
haftmann@25368
|
695 |
(map o apsnd) (fn (class, (ty, _)) => subst_class_typ [class] ty) operations;
|
wenzelm@25318
|
696 |
fun declare_const (c, _) =
|
wenzelm@25318
|
697 |
let val b = Sign.base_name c
|
wenzelm@25344
|
698 |
in Sign.intern_const thy b = c ? Variable.declare_const (b, c) end;
|
haftmann@25368
|
699 |
val unchecks = map (fn (c, (_, (ty, t))) => (t, Const (c, ty))) operations;
|
haftmann@25083
|
700 |
in
|
haftmann@25083
|
701 |
ctxt
|
haftmann@25368
|
702 |
|> fold declare_const local_constraints
|
haftmann@25368
|
703 |
|> fold (ProofContext.add_const_constraint o apsnd SOME) local_constraints
|
haftmann@25368
|
704 |
|> ClassSyntax.put {
|
haftmann@25368
|
705 |
local_constraints = local_constraints,
|
haftmann@25368
|
706 |
global_constraints = global_constraints,
|
haftmann@25083
|
707 |
base_sort = base_sort,
|
haftmann@25368
|
708 |
operations = (map o apsnd) snd operations,
|
haftmann@25195
|
709 |
unchecks = unchecks,
|
haftmann@25083
|
710 |
passed = false
|
haftmann@25368
|
711 |
}
|
haftmann@25083
|
712 |
end;
|
haftmann@25083
|
713 |
|
haftmann@25083
|
714 |
fun refresh_syntax class ctxt =
|
haftmann@25002
|
715 |
let
|
haftmann@25002
|
716 |
val thy = ProofContext.theory_of ctxt;
|
haftmann@25062
|
717 |
val base_sort = (#base_sort o the_class_data thy) class;
|
wenzelm@25344
|
718 |
in synchronize_syntax [class] base_sort ctxt end;
|
haftmann@24914
|
719 |
|
haftmann@25368
|
720 |
val mark_passed = ClassSyntax.map
|
haftmann@25368
|
721 |
(fn { local_constraints, global_constraints, base_sort, operations, unchecks, passed } =>
|
haftmann@25368
|
722 |
{ local_constraints = local_constraints, global_constraints = global_constraints,
|
haftmann@25368
|
723 |
base_sort = base_sort, operations = operations, unchecks = unchecks, passed = true });
|
haftmann@25083
|
724 |
|
haftmann@25083
|
725 |
fun sort_term_check ts ctxt =
|
haftmann@24748
|
726 |
let
|
haftmann@25368
|
727 |
val { local_constraints, global_constraints, base_sort, operations, passed, ... } =
|
haftmann@25368
|
728 |
ClassSyntax.get ctxt;
|
haftmann@25368
|
729 |
fun check_improve (Const (c, ty)) = (case AList.lookup (op =) local_constraints c
|
haftmann@25368
|
730 |
of SOME ty0 => (case try (Type.raw_match (ty0, ty)) Vartab.empty
|
haftmann@25368
|
731 |
of SOME tyenv => (case Vartab.lookup tyenv (Name.aT, 0)
|
haftmann@25368
|
732 |
of SOME (_, TVar (tvar as (vi, _))) =>
|
haftmann@25368
|
733 |
if TypeInfer.is_param vi then cons tvar else I
|
haftmann@25368
|
734 |
| _ => I)
|
haftmann@25368
|
735 |
| NONE => I)
|
haftmann@25368
|
736 |
| NONE => I)
|
haftmann@25368
|
737 |
| check_improve _ = I;
|
haftmann@25368
|
738 |
val improvements = (fold o fold_aterms) check_improve ts [];
|
haftmann@25368
|
739 |
val ts' = (map o map_types o map_atyps) (fn ty as TVar tvar =>
|
haftmann@25368
|
740 |
if member (op =) improvements tvar
|
haftmann@25368
|
741 |
then TFree (Name.aT, base_sort) else ty | ty => ty) ts;
|
haftmann@25368
|
742 |
fun check t0 = Envir.expand_term (fn Const (c, ty) => (case AList.lookup (op =) operations c
|
haftmann@25368
|
743 |
of SOME (ty0, t) =>
|
haftmann@25368
|
744 |
if Type.typ_instance (ProofContext.tsig_of ctxt) (ty, ty0)
|
haftmann@25368
|
745 |
then SOME (ty0, check t) else NONE
|
haftmann@25368
|
746 |
| NONE => NONE)
|
haftmann@25368
|
747 |
| _ => NONE) t0;
|
haftmann@25368
|
748 |
val ts'' = map check ts';
|
haftmann@25368
|
749 |
in if eq_list (op aconv) (ts, ts'') andalso passed then NONE
|
haftmann@25083
|
750 |
else
|
haftmann@25083
|
751 |
ctxt
|
haftmann@25368
|
752 |
|> fold (ProofContext.add_const_constraint o apsnd SOME) global_constraints
|
haftmann@25083
|
753 |
|> mark_passed
|
haftmann@25368
|
754 |
|> pair ts''
|
haftmann@25083
|
755 |
|> SOME
|
haftmann@25083
|
756 |
end;
|
haftmann@24748
|
757 |
|
haftmann@25083
|
758 |
fun sort_term_uncheck ts ctxt =
|
haftmann@25002
|
759 |
let
|
haftmann@25002
|
760 |
val thy = ProofContext.theory_of ctxt;
|
haftmann@25368
|
761 |
val unchecks = (#unchecks o ClassSyntax.get) ctxt;
|
haftmann@25462
|
762 |
val ts' = map (Pattern.rewrite_term thy unchecks []) ts;
|
wenzelm@25060
|
763 |
in if eq_list (op aconv) (ts, ts') then NONE else SOME (ts', ctxt) end;
|
haftmann@25002
|
764 |
|
wenzelm@25344
|
765 |
fun init_ctxt sups base_sort ctxt =
|
haftmann@25083
|
766 |
ctxt
|
haftmann@25083
|
767 |
|> Variable.declare_term
|
haftmann@25083
|
768 |
(Logic.mk_type (TFree (Name.aT, base_sort)))
|
wenzelm@25344
|
769 |
|> synchronize_syntax sups base_sort
|
haftmann@25083
|
770 |
|> Context.proof_map (
|
haftmann@25083
|
771 |
Syntax.add_term_check 0 "class" sort_term_check
|
haftmann@25103
|
772 |
#> Syntax.add_term_uncheck 0 "class" sort_term_uncheck)
|
haftmann@24901
|
773 |
|
haftmann@25311
|
774 |
fun init class thy =
|
haftmann@25311
|
775 |
thy
|
haftmann@25311
|
776 |
|> Locale.init class
|
wenzelm@25344
|
777 |
|> init_ctxt [class] ((#base_sort o the_class_data thy) class);
|
haftmann@24914
|
778 |
|
haftmann@24748
|
779 |
|
haftmann@24589
|
780 |
(* class definition *)
|
haftmann@24218
|
781 |
|
haftmann@24218
|
782 |
local
|
haftmann@24218
|
783 |
|
haftmann@24748
|
784 |
fun gen_class_spec prep_class prep_expr process_expr thy raw_supclasses raw_includes_elems =
|
haftmann@24218
|
785 |
let
|
haftmann@24748
|
786 |
val supclasses = map (prep_class thy) raw_supclasses;
|
haftmann@25209
|
787 |
val sups = filter (is_class thy) supclasses;
|
haftmann@25209
|
788 |
fun the_base_sort class = lookup_class_data thy class
|
haftmann@25209
|
789 |
|> Option.map #base_sort
|
haftmann@25209
|
790 |
|> the_default [class];
|
haftmann@25209
|
791 |
val base_sort = Sign.minimize_sort thy (maps the_base_sort supclasses);
|
haftmann@24748
|
792 |
val supsort = Sign.minimize_sort thy supclasses;
|
haftmann@25038
|
793 |
val suplocales = map Locale.Locale sups;
|
haftmann@24748
|
794 |
val (raw_elems, includes) = fold_rev (fn Locale.Elem e => apfst (cons e)
|
haftmann@24748
|
795 |
| Locale.Expr i => apsnd (cons (prep_expr thy i))) raw_includes_elems ([], []);
|
haftmann@24748
|
796 |
val supexpr = Locale.Merge suplocales;
|
haftmann@24748
|
797 |
val supparams = (map fst o Locale.parameters_of_expr thy) supexpr;
|
haftmann@25002
|
798 |
val supconsts = AList.make (the o AList.lookup (op =) (these_params thy sups))
|
haftmann@24748
|
799 |
(map fst supparams);
|
haftmann@24748
|
800 |
val mergeexpr = Locale.Merge (suplocales @ includes);
|
haftmann@24748
|
801 |
val constrain = Element.Constrains ((map o apsnd o map_atyps)
|
wenzelm@24847
|
802 |
(fn TFree (_, sort) => TFree (Name.aT, sort)) supparams);
|
haftmann@24748
|
803 |
in
|
haftmann@24748
|
804 |
ProofContext.init thy
|
haftmann@24748
|
805 |
|> Locale.cert_expr supexpr [constrain]
|
haftmann@24748
|
806 |
|> snd
|
wenzelm@25344
|
807 |
|> init_ctxt sups base_sort
|
haftmann@24748
|
808 |
|> process_expr Locale.empty raw_elems
|
haftmann@24748
|
809 |
|> fst
|
haftmann@25062
|
810 |
|> (fn elems => ((((sups, supconsts), (supsort, base_sort, mergeexpr)),
|
haftmann@24748
|
811 |
(*FIXME*) if null includes then constrain :: elems else elems)))
|
haftmann@24748
|
812 |
end;
|
haftmann@24748
|
813 |
|
haftmann@24748
|
814 |
val read_class_spec = gen_class_spec Sign.intern_class Locale.intern_expr Locale.read_expr;
|
haftmann@24748
|
815 |
val check_class_spec = gen_class_spec (K I) (K I) Locale.cert_expr;
|
haftmann@24748
|
816 |
|
wenzelm@24968
|
817 |
fun define_class_params (name, raw_superclasses) raw_consts raw_dep_axioms other_consts thy =
|
wenzelm@24968
|
818 |
let
|
wenzelm@24968
|
819 |
val superclasses = map (Sign.certify_class thy) raw_superclasses;
|
wenzelm@24968
|
820 |
val consts = (map o apfst o apsnd) (Sign.certify_typ thy) raw_consts;
|
haftmann@25083
|
821 |
fun add_const ((c, ty), syn) =
|
haftmann@25083
|
822 |
Sign.declare_const [] (c, Type.strip_sorts ty, syn) #>> Term.dest_Const;
|
wenzelm@24968
|
823 |
fun mk_axioms cs thy =
|
wenzelm@24968
|
824 |
raw_dep_axioms thy cs
|
wenzelm@24968
|
825 |
|> (map o apsnd o map) (Sign.cert_prop thy)
|
wenzelm@24968
|
826 |
|> rpair thy;
|
haftmann@25002
|
827 |
fun constrain_typs class = (map o apsnd o Term.map_type_tfree)
|
haftmann@25002
|
828 |
(fn (v, _) => TFree (v, [class]))
|
wenzelm@24968
|
829 |
in
|
wenzelm@24968
|
830 |
thy
|
wenzelm@24968
|
831 |
|> Sign.add_path (Logic.const_of_class name)
|
wenzelm@24968
|
832 |
|> fold_map add_const consts
|
wenzelm@24968
|
833 |
||> Sign.restore_naming thy
|
wenzelm@24968
|
834 |
|-> (fn cs => mk_axioms cs
|
wenzelm@24968
|
835 |
#-> (fn axioms_prop => AxClass.define_class (name, superclasses)
|
wenzelm@24968
|
836 |
(map fst cs @ other_consts) axioms_prop
|
haftmann@25002
|
837 |
#-> (fn class => `(fn _ => constrain_typs class cs)
|
haftmann@25002
|
838 |
#-> (fn cs' => `(fn thy => AxClass.get_info thy class)
|
haftmann@25002
|
839 |
#-> (fn {axioms, ...} => fold (Sign.add_const_constraint o apsnd SOME) cs'
|
haftmann@25002
|
840 |
#> pair (class, (cs', axioms)))))))
|
wenzelm@24968
|
841 |
end;
|
wenzelm@24968
|
842 |
|
haftmann@25002
|
843 |
fun gen_class prep_spec prep_param bname
|
haftmann@24748
|
844 |
raw_supclasses raw_includes_elems raw_other_consts thy =
|
haftmann@24748
|
845 |
let
|
haftmann@25038
|
846 |
val class = Sign.full_name thy bname;
|
haftmann@25062
|
847 |
val (((sups, supconsts), (supsort, base_sort, mergeexpr)), elems_syn) =
|
haftmann@24748
|
848 |
prep_spec thy raw_supclasses raw_includes_elems;
|
wenzelm@24968
|
849 |
val other_consts = map (tap (Sign.the_const_type thy) o prep_param thy) raw_other_consts;
|
haftmann@25163
|
850 |
fun mk_inst class cs =
|
haftmann@25163
|
851 |
(map o apsnd o Term.map_type_tfree) (fn (v, _) => TFree (v, [class])) cs;
|
haftmann@25062
|
852 |
fun fork_syntax (Element.Fixes xs) =
|
haftmann@25062
|
853 |
fold_map (fn (c, ty, syn) => cons (c, syn) #> pair (c, ty, NoSyn)) xs
|
haftmann@25062
|
854 |
#>> Element.Fixes
|
haftmann@25062
|
855 |
| fork_syntax x = pair x;
|
haftmann@25062
|
856 |
val (elems, global_syn) = fold_map fork_syntax elems_syn [];
|
wenzelm@25326
|
857 |
fun globalize (c, ty) =
|
haftmann@25062
|
858 |
((c, Term.map_type_tfree (K (TFree (Name.aT, base_sort))) ty),
|
haftmann@25062
|
859 |
(the_default NoSyn o AList.lookup (op =) global_syn) c);
|
haftmann@25038
|
860 |
fun extract_params thy =
|
haftmann@24218
|
861 |
let
|
haftmann@25062
|
862 |
val params = map fst (Locale.parameters_of thy class);
|
haftmann@24218
|
863 |
in
|
haftmann@25062
|
864 |
(params, (map globalize o snd o chop (length supconsts)) params)
|
haftmann@24218
|
865 |
end;
|
haftmann@25038
|
866 |
fun extract_assumes params thy cs =
|
haftmann@24218
|
867 |
let
|
haftmann@24218
|
868 |
val consts = supconsts @ (map (fst o fst) params ~~ cs);
|
haftmann@24218
|
869 |
fun subst (Free (c, ty)) =
|
haftmann@24218
|
870 |
Const ((fst o the o AList.lookup (op =) consts) c, ty)
|
haftmann@24218
|
871 |
| subst t = t;
|
haftmann@24218
|
872 |
fun prep_asm ((name, atts), ts) =
|
wenzelm@25024
|
873 |
((Sign.base_name name, map (Attrib.attribute_i thy) atts),
|
haftmann@24589
|
874 |
(map o map_aterms) subst ts);
|
haftmann@24218
|
875 |
in
|
haftmann@25038
|
876 |
Locale.global_asms_of thy class
|
haftmann@24218
|
877 |
|> map prep_asm
|
haftmann@24218
|
878 |
end;
|
haftmann@24218
|
879 |
in
|
haftmann@24218
|
880 |
thy
|
haftmann@24748
|
881 |
|> Locale.add_locale_i (SOME "") bname mergeexpr elems
|
haftmann@25038
|
882 |
|> snd
|
haftmann@25311
|
883 |
|> ProofContext.theory_of
|
haftmann@25311
|
884 |
|> `extract_params
|
haftmann@25311
|
885 |
|-> (fn (all_params, params) =>
|
wenzelm@24968
|
886 |
define_class_params (bname, supsort) params
|
haftmann@25038
|
887 |
(extract_assumes params) other_consts
|
haftmann@25038
|
888 |
#-> (fn (_, (consts, axioms)) =>
|
haftmann@25038
|
889 |
`(fn thy => class_intro thy class sups)
|
haftmann@24218
|
890 |
#-> (fn class_intro =>
|
haftmann@25062
|
891 |
PureThy.note_thmss_qualified "" (NameSpace.append class classN)
|
haftmann@25062
|
892 |
[((introN, []), [([class_intro], [])])]
|
haftmann@25062
|
893 |
#-> (fn [(_, [class_intro])] =>
|
haftmann@25038
|
894 |
add_class_data ((class, sups),
|
haftmann@25062
|
895 |
(map fst params ~~ consts, base_sort,
|
haftmann@25163
|
896 |
mk_inst class (map snd supconsts @ consts),
|
haftmann@25062
|
897 |
calculate_morphism class (supconsts @ (map (fst o fst) params ~~ consts)), class_intro))
|
haftmann@25038
|
898 |
#> class_interpretation class axioms []
|
haftmann@25311
|
899 |
))))
|
haftmann@25268
|
900 |
|> init class
|
haftmann@25038
|
901 |
|> pair class
|
haftmann@24218
|
902 |
end;
|
haftmann@24218
|
903 |
|
wenzelm@25326
|
904 |
fun read_const thy = #1 o Term.dest_Const o ProofContext.read_const (ProofContext.init thy);
|
wenzelm@25326
|
905 |
|
haftmann@24218
|
906 |
in
|
haftmann@24218
|
907 |
|
wenzelm@25326
|
908 |
val class_cmd = gen_class read_class_spec read_const;
|
haftmann@24748
|
909 |
val class = gen_class check_class_spec (K I);
|
haftmann@24218
|
910 |
|
haftmann@24218
|
911 |
end; (*local*)
|
haftmann@24218
|
912 |
|
haftmann@24218
|
913 |
|
haftmann@24589
|
914 |
(* definition in class target *)
|
haftmann@24218
|
915 |
|
haftmann@25096
|
916 |
fun add_logical_const class pos ((c, mx), dict) thy =
|
haftmann@24218
|
917 |
let
|
wenzelm@25024
|
918 |
val prfx = class_prefix class;
|
wenzelm@25024
|
919 |
val thy' = thy |> Sign.add_path prfx;
|
haftmann@25062
|
920 |
val phi = morphism thy' class;
|
wenzelm@25024
|
921 |
|
haftmann@25062
|
922 |
val c' = Sign.full_name thy' c;
|
haftmann@25239
|
923 |
val dict' = Morphism.term phi dict;
|
haftmann@25239
|
924 |
val dict_def = map_types Logic.unvarifyT dict';
|
haftmann@25239
|
925 |
val ty' = Term.fastype_of dict_def;
|
haftmann@25083
|
926 |
val ty'' = Type.strip_sorts ty';
|
haftmann@25239
|
927 |
val def_eq = Logic.mk_equals (Const (c', ty'), dict_def);
|
haftmann@24218
|
928 |
in
|
wenzelm@25024
|
929 |
thy'
|
haftmann@25096
|
930 |
|> Sign.declare_const pos (c, ty'', mx) |> snd
|
wenzelm@25318
|
931 |
|> Thm.add_def false (c, def_eq) (* FIXME PureThy.add_defs_i *)
|
haftmann@25062
|
932 |
|>> Thm.symmetric
|
haftmann@25083
|
933 |
|-> (fn def => class_interpretation class [def] [Thm.prop_of def]
|
haftmann@25368
|
934 |
#> register_operation class (c', (dict', SOME (Thm.varifyT def))))
|
haftmann@24218
|
935 |
|> Sign.restore_naming thy
|
haftmann@25083
|
936 |
|> Sign.add_const_constraint (c', SOME ty')
|
haftmann@24218
|
937 |
end;
|
haftmann@24218
|
938 |
|
haftmann@24901
|
939 |
|
haftmann@24901
|
940 |
(* abbreviation in class target *)
|
haftmann@24901
|
941 |
|
haftmann@25103
|
942 |
fun add_syntactic_const class prmode pos ((c, mx), rhs) thy =
|
haftmann@24836
|
943 |
let
|
wenzelm@25024
|
944 |
val prfx = class_prefix class;
|
haftmann@25096
|
945 |
val thy' = thy |> Sign.add_path prfx;
|
haftmann@25062
|
946 |
val phi = morphism thy class;
|
haftmann@25062
|
947 |
|
haftmann@25096
|
948 |
val c' = Sign.full_name thy' c;
|
haftmann@25146
|
949 |
val rews = map (Logic.dest_equals o Thm.prop_of) (these_defs thy' [class])
|
haftmann@25146
|
950 |
val rhs' = (Pattern.rewrite_term thy rews [] o Morphism.term phi) rhs;
|
haftmann@25239
|
951 |
val ty' = Logic.unvarifyT (Term.fastype_of rhs');
|
haftmann@24836
|
952 |
in
|
haftmann@25096
|
953 |
thy'
|
haftmann@25146
|
954 |
|> Sign.add_abbrev (#1 prmode) pos (c, map_types Type.strip_sorts rhs') |> snd
|
haftmann@25096
|
955 |
|> Sign.add_const_constraint (c', SOME ty')
|
wenzelm@25024
|
956 |
|> Sign.notation true prmode [(Const (c', ty'), mx)]
|
haftmann@25368
|
957 |
|> register_operation class (c', (rhs', NONE))
|
haftmann@25096
|
958 |
|> Sign.restore_naming thy
|
haftmann@24836
|
959 |
end;
|
haftmann@24836
|
960 |
|
haftmann@25462
|
961 |
|
haftmann@25462
|
962 |
(** instantiation target **)
|
haftmann@25462
|
963 |
|
haftmann@25462
|
964 |
(* bookkeeping *)
|
haftmann@25462
|
965 |
|
haftmann@25462
|
966 |
datatype instantiation = Instantiation of {
|
haftmann@25462
|
967 |
arities: arity list,
|
haftmann@25462
|
968 |
params: ((string * string) * (string * typ)) list
|
haftmann@25462
|
969 |
}
|
haftmann@25462
|
970 |
|
haftmann@25462
|
971 |
structure Instantiation = ProofDataFun
|
haftmann@25462
|
972 |
(
|
haftmann@25462
|
973 |
type T = instantiation
|
haftmann@25462
|
974 |
fun init _ = Instantiation { arities = [], params = [] };
|
haftmann@25462
|
975 |
);
|
haftmann@25462
|
976 |
|
haftmann@25462
|
977 |
fun mk_instantiation (arities, params) = Instantiation {
|
haftmann@25462
|
978 |
arities = arities, params = params
|
haftmann@25462
|
979 |
};
|
haftmann@25462
|
980 |
fun map_instantiation f (Instantiation { arities, params }) =
|
haftmann@25462
|
981 |
mk_instantiation (f (arities, params));
|
haftmann@25462
|
982 |
|
haftmann@25462
|
983 |
fun the_instantiation ctxt = case Instantiation.get ctxt
|
haftmann@25462
|
984 |
of Instantiation { arities = [], ... } => error "No instantiation target"
|
haftmann@25462
|
985 |
| Instantiation data => data;
|
haftmann@25462
|
986 |
|
haftmann@25462
|
987 |
fun init_instantiation arities ctxt =
|
haftmann@25462
|
988 |
let
|
haftmann@25462
|
989 |
val thy = ProofContext.theory_of ctxt;
|
haftmann@25462
|
990 |
val _ = if null arities then error "At least one arity must be given" else ();
|
haftmann@25462
|
991 |
val _ = case (duplicates (op =) o map #1) arities
|
haftmann@25462
|
992 |
of [] => ()
|
haftmann@25462
|
993 |
| dupl_tycos => error ("Type constructors occur more than once in arities: "
|
haftmann@25462
|
994 |
^ commas_quote dupl_tycos);
|
haftmann@25462
|
995 |
val ty_insts = map (fn (tyco, sorts, _) =>
|
haftmann@25462
|
996 |
(tyco, Type (tyco, map TFree (Name.names Name.context Name.aT sorts))))
|
haftmann@25462
|
997 |
arities;
|
haftmann@25462
|
998 |
val ty_inst = the o AList.lookup (op =) ty_insts;
|
haftmann@25462
|
999 |
fun type_name "*" = "prod"
|
haftmann@25462
|
1000 |
| type_name "+" = "sum"
|
haftmann@25462
|
1001 |
| type_name s = NameSpace.base s; (*FIXME*)
|
haftmann@25462
|
1002 |
fun get_param tyco sorts (param, (c, ty)) =
|
haftmann@25462
|
1003 |
((unoverload_const thy (c, ty), tyco),
|
haftmann@25462
|
1004 |
(param ^ "_" ^ type_name tyco,
|
haftmann@25462
|
1005 |
map_atyps (K (ty_inst tyco)) ty));
|
haftmann@25462
|
1006 |
fun get_params (tyco, sorts, sort) =
|
haftmann@25462
|
1007 |
map (get_param tyco sorts) (these_params thy sort)
|
haftmann@25462
|
1008 |
val params = maps get_params arities;
|
haftmann@25462
|
1009 |
in
|
haftmann@25462
|
1010 |
ctxt
|
haftmann@25462
|
1011 |
|> Instantiation.put (mk_instantiation (arities, params))
|
haftmann@25462
|
1012 |
|> fold (Variable.declare_term o Logic.mk_type o snd) ty_insts
|
haftmann@25462
|
1013 |
|> fold (Variable.declare_term o Free o snd) params
|
haftmann@25462
|
1014 |
end;
|
haftmann@25462
|
1015 |
|
haftmann@25462
|
1016 |
val instantiation_params = #params o the_instantiation;
|
haftmann@25462
|
1017 |
|
haftmann@25462
|
1018 |
fun instantiation_const ctxt v = instantiation_params ctxt
|
haftmann@25462
|
1019 |
|> find_first (fn (_, (v', _)) => v = v')
|
haftmann@25462
|
1020 |
|> Option.map (fst o fst);
|
haftmann@25462
|
1021 |
|
haftmann@25462
|
1022 |
|
haftmann@25462
|
1023 |
(* syntax *)
|
haftmann@25462
|
1024 |
|
haftmann@25462
|
1025 |
fun inst_term_check ts ctxt =
|
haftmann@25462
|
1026 |
let
|
haftmann@25462
|
1027 |
val params = instantiation_params ctxt;
|
haftmann@25462
|
1028 |
val tsig = ProofContext.tsig_of ctxt;
|
haftmann@25462
|
1029 |
val thy = ProofContext.theory_of ctxt;
|
haftmann@25462
|
1030 |
|
haftmann@25462
|
1031 |
fun check_improve (Const (c, ty)) = (case inst_tyco thy (c, ty)
|
haftmann@25462
|
1032 |
of SOME tyco => (case AList.lookup (op =) params (c, tyco)
|
haftmann@25462
|
1033 |
of SOME (_, ty') => Type.typ_match tsig (ty, ty')
|
haftmann@25462
|
1034 |
| NONE => I)
|
haftmann@25462
|
1035 |
| NONE => I)
|
haftmann@25462
|
1036 |
| check_improve _ = I;
|
haftmann@25462
|
1037 |
val improvement = (fold o fold_aterms) check_improve ts Vartab.empty;
|
haftmann@25462
|
1038 |
val ts' = (map o map_types) (Envir.typ_subst_TVars improvement) ts;
|
haftmann@25462
|
1039 |
val ts'' = (map o map_aterms) (fn t as Const (c, ty) => (case inst_tyco thy (c, ty)
|
haftmann@25462
|
1040 |
of SOME tyco => (case AList.lookup (op =) params (c, tyco)
|
haftmann@25462
|
1041 |
of SOME v_ty => Free v_ty
|
haftmann@25462
|
1042 |
| NONE => t)
|
haftmann@25462
|
1043 |
| NONE => t)
|
haftmann@25462
|
1044 |
| t => t) ts';
|
haftmann@25462
|
1045 |
in if eq_list (op aconv) (ts, ts'') then NONE else SOME (ts'', ctxt) end;
|
haftmann@25462
|
1046 |
|
haftmann@25462
|
1047 |
fun inst_term_uncheck ts ctxt =
|
haftmann@25462
|
1048 |
let
|
haftmann@25462
|
1049 |
val params = instantiation_params ctxt;
|
haftmann@25462
|
1050 |
val ts' = (map o map_aterms) (fn t as Free (v, ty) =>
|
haftmann@25462
|
1051 |
(case get_first (fn ((c, _), (v', _)) => if v = v' then SOME c else NONE) params
|
haftmann@25462
|
1052 |
of SOME c => Const (c, ty)
|
haftmann@25462
|
1053 |
| NONE => t)
|
haftmann@25462
|
1054 |
| t => t) ts;
|
haftmann@25462
|
1055 |
in if eq_list (op aconv) (ts, ts') then NONE else SOME (ts', ctxt) end;
|
haftmann@25462
|
1056 |
|
haftmann@25462
|
1057 |
|
haftmann@25462
|
1058 |
(* target *)
|
haftmann@25462
|
1059 |
|
haftmann@25462
|
1060 |
fun instantiation arities =
|
haftmann@25462
|
1061 |
ProofContext.init
|
haftmann@25462
|
1062 |
#> init_instantiation arities
|
haftmann@25462
|
1063 |
#> fold ProofContext.add_arity arities
|
haftmann@25462
|
1064 |
#> Context.proof_map (
|
haftmann@25462
|
1065 |
Syntax.add_term_check 0 "instance" inst_term_check
|
haftmann@25462
|
1066 |
#> Syntax.add_term_uncheck 0 "instance" inst_term_uncheck);
|
haftmann@25462
|
1067 |
|
haftmann@25462
|
1068 |
fun gen_proof_instantiation do_proof after_qed lthy =
|
haftmann@25462
|
1069 |
let
|
haftmann@25462
|
1070 |
(*FIXME should work on fresh context but continue local theory afterwards*)
|
haftmann@25462
|
1071 |
val ctxt = LocalTheory.target_of lthy;
|
haftmann@25462
|
1072 |
val arities = (#arities o the_instantiation) ctxt;
|
haftmann@25462
|
1073 |
val arities_proof = maps
|
haftmann@25462
|
1074 |
(Logic.mk_arities o Sign.cert_arity (ProofContext.theory_of ctxt)) arities;
|
haftmann@25462
|
1075 |
fun after_qed' results =
|
haftmann@25462
|
1076 |
LocalTheory.theory (fold (AxClass.add_arity o Thm.varifyT) results)
|
haftmann@25462
|
1077 |
#> after_qed;
|
haftmann@25462
|
1078 |
in
|
haftmann@25462
|
1079 |
lthy
|
haftmann@25462
|
1080 |
|> do_proof after_qed' arities_proof
|
haftmann@25462
|
1081 |
end;
|
haftmann@25462
|
1082 |
|
haftmann@25462
|
1083 |
val proof_instantiation = gen_proof_instantiation (fn after_qed => fn ts =>
|
haftmann@25462
|
1084 |
Proof.theorem_i NONE (after_qed o map the_single) (map (fn t => [(t, [])]) ts));
|
haftmann@25462
|
1085 |
|
haftmann@25462
|
1086 |
fun prove_instantiation tac = gen_proof_instantiation (fn after_qed =>
|
haftmann@25462
|
1087 |
fn ts => fn lthy => after_qed (Goal.prove_multi lthy [] [] ts
|
haftmann@25462
|
1088 |
(fn {context, ...} => tac context)) lthy) I;
|
haftmann@25462
|
1089 |
|
haftmann@25462
|
1090 |
fun conclude_instantiation lthy =
|
haftmann@25462
|
1091 |
let
|
haftmann@25462
|
1092 |
val arities = (#arities o the_instantiation) lthy;
|
haftmann@25462
|
1093 |
val thy = ProofContext.theory_of lthy;
|
haftmann@25462
|
1094 |
(*val _ = map (fn (tyco, sorts, sort) =>
|
haftmann@25462
|
1095 |
if Sign.of_sort thy
|
haftmann@25462
|
1096 |
(Type (tyco, map TFree (Name.names Name.context Name.aT sorts)), sort)
|
haftmann@25462
|
1097 |
then () else error ("Missing instance proof for type " ^ quote (Sign.extern_type thy tyco)))
|
haftmann@25462
|
1098 |
arities; FIXME activate when old instance command is gone*)
|
haftmann@25462
|
1099 |
val params_of = maps (these o try (#params o AxClass.get_info thy))
|
haftmann@25462
|
1100 |
o Sign.complete_sort thy;
|
haftmann@25462
|
1101 |
val missing_params = arities
|
haftmann@25462
|
1102 |
|> maps (fn (tyco, _, sort) => params_of sort |> map (rpair tyco))
|
haftmann@25462
|
1103 |
|> filter_out (can (inst thy) o apfst fst);
|
haftmann@25462
|
1104 |
fun declare_missing ((c, ty), tyco) thy =
|
haftmann@25462
|
1105 |
let
|
haftmann@25462
|
1106 |
val SOME class = AxClass.class_of_param thy c;
|
haftmann@25462
|
1107 |
val name_inst = NameSpace.base class ^ "_" ^ NameSpace.base tyco ^ "_inst";
|
haftmann@25462
|
1108 |
val vs = Name.names Name.context Name.aT (replicate (Sign.arity_number thy tyco) []);
|
haftmann@25462
|
1109 |
val ty' = map_atyps (fn _ => Type (tyco, map TFree vs)) ty;
|
haftmann@25462
|
1110 |
val c' = NameSpace.base c;
|
haftmann@25462
|
1111 |
in
|
haftmann@25462
|
1112 |
thy
|
haftmann@25462
|
1113 |
|> Sign.sticky_prefix name_inst
|
haftmann@25462
|
1114 |
|> Sign.no_base_names
|
haftmann@25462
|
1115 |
|> Sign.declare_const [] (c', ty', NoSyn)
|
haftmann@25462
|
1116 |
|-> (fn const' as Const (c'', _) => Thm.add_def true
|
haftmann@25462
|
1117 |
(Thm.def_name c', Logic.mk_equals (const', Const (c, ty')))
|
haftmann@25462
|
1118 |
#>> Thm.varifyT
|
haftmann@25462
|
1119 |
#-> (fn thm => add_inst (c, tyco) (c'', Thm.symmetric thm)
|
haftmann@25462
|
1120 |
#> primitive_note Thm.internalK (c', thm)
|
haftmann@25462
|
1121 |
#> snd
|
haftmann@25462
|
1122 |
#> Sign.restore_naming thy))
|
haftmann@25462
|
1123 |
end;
|
haftmann@25462
|
1124 |
in
|
haftmann@25462
|
1125 |
lthy
|
haftmann@25462
|
1126 |
|> LocalTheory.theory (fold declare_missing missing_params)
|
haftmann@25462
|
1127 |
end;
|
haftmann@25462
|
1128 |
|
haftmann@25462
|
1129 |
val end_instantiation = conclude_instantiation #> LocalTheory.target_of;
|
haftmann@25462
|
1130 |
|
haftmann@24218
|
1131 |
end;
|