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@24218
|
10 |
val axclass_cmd: bstring * xstring list
|
haftmann@24589
|
11 |
-> ((bstring * Attrib.src list) * string list) list
|
haftmann@24589
|
12 |
-> theory -> class * theory
|
haftmann@24589
|
13 |
val classrel_cmd: xstring * xstring -> theory -> Proof.state
|
haftmann@24914
|
14 |
|
haftmann@25002
|
15 |
val class: bstring -> class list -> Element.context_i Locale.element list
|
haftmann@24218
|
16 |
-> string list -> theory -> string * Proof.context
|
haftmann@25002
|
17 |
val class_cmd: bstring -> xstring list -> Element.context Locale.element list
|
haftmann@24589
|
18 |
-> xstring list -> theory -> string * Proof.context
|
haftmann@25038
|
19 |
val is_class: theory -> class -> bool
|
haftmann@25002
|
20 |
val these_params: theory -> sort -> (string * (string * typ)) list
|
haftmann@25083
|
21 |
val init: class -> Proof.context -> Proof.context
|
haftmann@25083
|
22 |
val add_logical_const: string -> (string * mixfix) * term
|
haftmann@25083
|
23 |
-> theory -> string * theory
|
haftmann@25083
|
24 |
val add_syntactic_const: string -> Syntax.mode -> (string * mixfix) * term
|
haftmann@25083
|
25 |
-> theory -> string * theory
|
haftmann@25083
|
26 |
val refresh_syntax: class -> Proof.context -> Proof.context
|
haftmann@24589
|
27 |
val intro_classes_tac: thm list -> tactic
|
haftmann@24589
|
28 |
val default_intro_classes_tac: thm list -> tactic
|
haftmann@24589
|
29 |
val print_classes: theory -> unit
|
haftmann@25002
|
30 |
val uncheck: bool ref
|
haftmann@24423
|
31 |
|
haftmann@24589
|
32 |
val instance_arity: (theory -> theory) -> arity list -> theory -> Proof.state
|
haftmann@24589
|
33 |
val instance: arity list -> ((bstring * Attrib.src list) * term) list
|
haftmann@24423
|
34 |
-> (thm list -> theory -> theory)
|
haftmann@24218
|
35 |
-> theory -> Proof.state
|
haftmann@24589
|
36 |
val instance_cmd: (bstring * xstring list * xstring) list
|
haftmann@24589
|
37 |
-> ((bstring * Attrib.src list) * xstring) list
|
haftmann@24423
|
38 |
-> (thm list -> theory -> theory)
|
haftmann@24218
|
39 |
-> theory -> Proof.state
|
haftmann@24589
|
40 |
val prove_instance: tactic -> arity list
|
haftmann@24218
|
41 |
-> ((bstring * Attrib.src list) * term) list
|
haftmann@24423
|
42 |
-> theory -> thm list * theory
|
haftmann@24423
|
43 |
val unoverload: theory -> conv
|
haftmann@24423
|
44 |
val overload: theory -> conv
|
haftmann@24423
|
45 |
val unoverload_const: theory -> string * typ -> string
|
haftmann@24423
|
46 |
val inst_const: theory -> string * string -> string
|
haftmann@24423
|
47 |
val param_const: theory -> string -> (string * string) option
|
haftmann@24218
|
48 |
end;
|
haftmann@24218
|
49 |
|
haftmann@24218
|
50 |
structure Class : CLASS =
|
haftmann@24218
|
51 |
struct
|
haftmann@24218
|
52 |
|
haftmann@24218
|
53 |
(** auxiliary **)
|
haftmann@24218
|
54 |
|
haftmann@25062
|
55 |
val classN = "class";
|
haftmann@25062
|
56 |
val introN = "intro";
|
haftmann@25062
|
57 |
|
haftmann@25002
|
58 |
fun prove_interpretation tac prfx_atts expr inst =
|
haftmann@25002
|
59 |
Locale.interpretation_i I prfx_atts expr inst
|
haftmann@24589
|
60 |
#> Proof.global_terminal_proof
|
haftmann@24589
|
61 |
(Method.Basic (K (Method.SIMPLE_METHOD tac), Position.none), NONE)
|
haftmann@24589
|
62 |
#> ProofContext.theory_of;
|
haftmann@24589
|
63 |
|
wenzelm@25020
|
64 |
fun OF_LAST thm1 thm2 = thm1 RSN (Thm.nprems_of thm2, thm2);
|
haftmann@24589
|
65 |
|
haftmann@24589
|
66 |
fun strip_all_ofclass thy sort =
|
haftmann@24589
|
67 |
let
|
wenzelm@24847
|
68 |
val typ = TVar ((Name.aT, 0), sort);
|
haftmann@24589
|
69 |
fun prem_inclass t =
|
haftmann@24589
|
70 |
case Logic.strip_imp_prems t
|
haftmann@24589
|
71 |
of ofcls :: _ => try Logic.dest_inclass ofcls
|
haftmann@24589
|
72 |
| [] => NONE;
|
haftmann@24589
|
73 |
fun strip_ofclass class thm =
|
haftmann@24589
|
74 |
thm OF (fst o AxClass.of_sort thy (typ, [class])) AxClass.cache;
|
haftmann@24589
|
75 |
fun strip thm = case (prem_inclass o Thm.prop_of) thm
|
haftmann@24589
|
76 |
of SOME (_, class) => thm |> strip_ofclass class |> strip
|
haftmann@24589
|
77 |
| NONE => thm;
|
haftmann@24589
|
78 |
in strip end;
|
haftmann@24589
|
79 |
|
haftmann@25038
|
80 |
fun get_remove_global_constraint c thy =
|
haftmann@25038
|
81 |
let
|
haftmann@25038
|
82 |
val ty = Sign.the_const_constraint thy c;
|
haftmann@25038
|
83 |
in
|
haftmann@25038
|
84 |
thy
|
haftmann@25038
|
85 |
|> Sign.add_const_constraint (c, NONE)
|
haftmann@25038
|
86 |
|> pair (c, Logic.unvarifyT ty)
|
haftmann@25038
|
87 |
end;
|
haftmann@25038
|
88 |
|
haftmann@24589
|
89 |
|
haftmann@24589
|
90 |
(** axclass command **)
|
haftmann@24589
|
91 |
|
haftmann@24218
|
92 |
fun axclass_cmd (class, raw_superclasses) raw_specs thy =
|
haftmann@24218
|
93 |
let
|
haftmann@24218
|
94 |
val ctxt = ProofContext.init thy;
|
haftmann@24218
|
95 |
val superclasses = map (Sign.read_class thy) raw_superclasses;
|
haftmann@24589
|
96 |
val name_atts = map ((apsnd o map) (Attrib.attribute thy) o fst)
|
haftmann@24589
|
97 |
raw_specs;
|
haftmann@24589
|
98 |
val axiomss = ProofContext.read_propp (ctxt, map (map (rpair []) o snd)
|
haftmann@24589
|
99 |
raw_specs)
|
haftmann@24218
|
100 |
|> snd
|
haftmann@24218
|
101 |
|> (map o map) fst;
|
haftmann@24589
|
102 |
in
|
haftmann@24589
|
103 |
AxClass.define_class (class, superclasses) []
|
haftmann@24589
|
104 |
(name_atts ~~ axiomss) thy
|
haftmann@24589
|
105 |
end;
|
haftmann@24218
|
106 |
|
haftmann@24218
|
107 |
local
|
haftmann@24218
|
108 |
|
haftmann@24218
|
109 |
fun gen_instance mk_prop add_thm after_qed insts thy =
|
haftmann@24218
|
110 |
let
|
haftmann@24218
|
111 |
fun after_qed' results =
|
haftmann@24218
|
112 |
ProofContext.theory ((fold o fold) add_thm results #> after_qed);
|
haftmann@24218
|
113 |
in
|
haftmann@24218
|
114 |
thy
|
haftmann@24218
|
115 |
|> ProofContext.init
|
haftmann@24589
|
116 |
|> Proof.theorem_i NONE after_qed' ((map (fn t => [(t, [])])
|
haftmann@24589
|
117 |
o maps (mk_prop thy)) insts)
|
haftmann@24218
|
118 |
end;
|
haftmann@24218
|
119 |
|
haftmann@24218
|
120 |
in
|
haftmann@24218
|
121 |
|
haftmann@24589
|
122 |
val instance_arity =
|
haftmann@24218
|
123 |
gen_instance (Logic.mk_arities oo Sign.cert_arity) AxClass.add_arity;
|
haftmann@24589
|
124 |
val classrel =
|
haftmann@24218
|
125 |
gen_instance (single oo (Logic.mk_classrel oo AxClass.cert_classrel))
|
haftmann@24218
|
126 |
AxClass.add_classrel I o single;
|
haftmann@24589
|
127 |
val classrel_cmd =
|
haftmann@24589
|
128 |
gen_instance (single oo (Logic.mk_classrel oo AxClass.read_classrel))
|
haftmann@24589
|
129 |
AxClass.add_classrel I o single;
|
haftmann@24218
|
130 |
|
haftmann@24218
|
131 |
end; (*local*)
|
haftmann@24218
|
132 |
|
haftmann@24218
|
133 |
|
haftmann@24589
|
134 |
(** explicit constants for overloaded definitions **)
|
haftmann@24304
|
135 |
|
haftmann@24304
|
136 |
structure InstData = TheoryDataFun
|
haftmann@24304
|
137 |
(
|
haftmann@24423
|
138 |
type T = (string * thm) Symtab.table Symtab.table * (string * string) Symtab.table;
|
haftmann@24423
|
139 |
(*constant name ~> type constructor ~> (constant name, equation),
|
haftmann@24423
|
140 |
constant name ~> (constant name, type constructor)*)
|
haftmann@24423
|
141 |
val empty = (Symtab.empty, Symtab.empty);
|
haftmann@24304
|
142 |
val copy = I;
|
haftmann@24304
|
143 |
val extend = I;
|
haftmann@24423
|
144 |
fun merge _ ((taba1, tabb1), (taba2, tabb2)) =
|
haftmann@24423
|
145 |
(Symtab.join (K (Symtab.merge (K true))) (taba1, taba2),
|
haftmann@24423
|
146 |
Symtab.merge (K true) (tabb1, tabb2));
|
haftmann@24304
|
147 |
);
|
haftmann@24304
|
148 |
|
wenzelm@25020
|
149 |
fun inst_thms thy = (Symtab.fold (Symtab.fold (cons o snd o snd) o snd) o fst)
|
haftmann@24589
|
150 |
(InstData.get thy) [];
|
haftmann@24589
|
151 |
fun add_inst (c, tyco) inst = (InstData.map o apfst
|
haftmann@24589
|
152 |
o Symtab.map_default (c, Symtab.empty)) (Symtab.update_new (tyco, inst))
|
haftmann@24423
|
153 |
#> (InstData.map o apsnd) (Symtab.update_new (fst inst, (c, tyco)));
|
haftmann@24304
|
154 |
|
wenzelm@25020
|
155 |
fun unoverload thy = MetaSimplifier.rewrite true (inst_thms thy);
|
wenzelm@25020
|
156 |
fun overload thy = MetaSimplifier.rewrite true (map Thm.symmetric (inst_thms thy));
|
haftmann@24304
|
157 |
|
haftmann@24304
|
158 |
fun inst_const thy (c, tyco) =
|
haftmann@24423
|
159 |
(fst o the o Symtab.lookup ((the o Symtab.lookup (fst (InstData.get thy))) c)) tyco;
|
haftmann@24423
|
160 |
fun unoverload_const thy (c_ty as (c, _)) =
|
haftmann@24423
|
161 |
case AxClass.class_of_param thy c
|
haftmann@24423
|
162 |
of SOME class => (case Sign.const_typargs thy c_ty
|
haftmann@24589
|
163 |
of [Type (tyco, _)] => (case Symtab.lookup
|
haftmann@24589
|
164 |
((the o Symtab.lookup (fst (InstData.get thy))) c) tyco
|
haftmann@24423
|
165 |
of SOME (c, _) => c
|
haftmann@24423
|
166 |
| NONE => c)
|
haftmann@24423
|
167 |
| [_] => c)
|
haftmann@24423
|
168 |
| NONE => c;
|
haftmann@24423
|
169 |
|
haftmann@24423
|
170 |
val param_const = Symtab.lookup o snd o InstData.get;
|
haftmann@24304
|
171 |
|
haftmann@24304
|
172 |
fun add_inst_def (class, tyco) (c, ty) thy =
|
haftmann@24304
|
173 |
let
|
wenzelm@25024
|
174 |
val tyco_base = Sign.base_name tyco;
|
wenzelm@25024
|
175 |
val name_inst = Sign.base_name class ^ "_" ^ tyco_base ^ "_inst";
|
wenzelm@25024
|
176 |
val c_inst_base = Sign.base_name c ^ "_" ^ tyco_base;
|
haftmann@24304
|
177 |
in
|
haftmann@24304
|
178 |
thy
|
haftmann@24304
|
179 |
|> Sign.sticky_prefix name_inst
|
wenzelm@25020
|
180 |
|> Sign.declare_const [] (c_inst_base, ty, NoSyn)
|
wenzelm@24968
|
181 |
|-> (fn const as Const (c_inst, _) =>
|
wenzelm@24968
|
182 |
PureThy.add_defs_i false
|
wenzelm@24968
|
183 |
[((Thm.def_name c_inst_base, Logic.mk_equals (const, Const (c, ty))), [])]
|
wenzelm@24968
|
184 |
#-> (fn [thm] => add_inst (c, tyco) (c_inst, Thm.symmetric thm)))
|
haftmann@24435
|
185 |
|> Sign.restore_naming thy
|
haftmann@24304
|
186 |
end;
|
haftmann@24304
|
187 |
|
haftmann@24304
|
188 |
fun add_inst_def' (class, tyco) (c, ty) thy =
|
haftmann@24423
|
189 |
if case Symtab.lookup (fst (InstData.get thy)) c
|
haftmann@24304
|
190 |
of NONE => true
|
haftmann@24304
|
191 |
| SOME tab => is_none (Symtab.lookup tab tyco)
|
haftmann@24304
|
192 |
then add_inst_def (class, tyco) (c, Logic.unvarifyT ty) thy
|
haftmann@24304
|
193 |
else thy;
|
haftmann@24304
|
194 |
|
haftmann@24304
|
195 |
fun add_def ((class, tyco), ((name, prop), atts)) thy =
|
haftmann@24304
|
196 |
let
|
haftmann@24589
|
197 |
val ((lhs as Const (c, ty), args), rhs) =
|
haftmann@24589
|
198 |
(apfst Term.strip_comb o Logic.dest_equals) prop;
|
haftmann@24701
|
199 |
fun (*add_inst' def ([], (Const (c_inst, ty))) =
|
haftmann@24423
|
200 |
if forall (fn TFree _ => true | _ => false) (Sign.const_typargs thy (c_inst, ty))
|
haftmann@24304
|
201 |
then add_inst (c, tyco) (c_inst, def)
|
haftmann@24304
|
202 |
else add_inst_def (class, tyco) (c, ty)
|
haftmann@24701
|
203 |
|*) add_inst' _ t = add_inst_def (class, tyco) (c, ty);
|
haftmann@24304
|
204 |
in
|
haftmann@24304
|
205 |
thy
|
wenzelm@25020
|
206 |
|> PureThy.add_defs_i true [((name, prop), map (Attrib.attribute_i thy) atts)]
|
haftmann@24304
|
207 |
|-> (fn [def] => add_inst' def (args, rhs) #> pair def)
|
haftmann@24304
|
208 |
end;
|
haftmann@24304
|
209 |
|
haftmann@24304
|
210 |
|
haftmann@24589
|
211 |
(** instances with implicit parameter handling **)
|
haftmann@24218
|
212 |
|
haftmann@24218
|
213 |
local
|
haftmann@24218
|
214 |
|
wenzelm@24707
|
215 |
fun gen_read_def thy prep_att parse_prop ((raw_name, raw_atts), raw_t) =
|
haftmann@24218
|
216 |
let
|
wenzelm@24707
|
217 |
val ctxt = ProofContext.init thy;
|
wenzelm@24707
|
218 |
val t = parse_prop ctxt raw_t |> Syntax.check_prop ctxt;
|
wenzelm@24981
|
219 |
val ((c, ty), _) = Sign.cert_def ctxt t;
|
haftmann@24218
|
220 |
val atts = map (prep_att thy) raw_atts;
|
haftmann@24218
|
221 |
val insts = Consts.typargs (Sign.consts_of thy) (c, ty);
|
haftmann@24218
|
222 |
val name = case raw_name
|
haftmann@24218
|
223 |
of "" => NONE
|
haftmann@24218
|
224 |
| _ => SOME raw_name;
|
haftmann@24218
|
225 |
in (c, (insts, ((name, t), atts))) end;
|
haftmann@24218
|
226 |
|
wenzelm@24707
|
227 |
fun read_def_cmd thy = gen_read_def thy Attrib.intern_src Syntax.parse_prop;
|
haftmann@24218
|
228 |
fun read_def thy = gen_read_def thy (K I) (K I);
|
haftmann@24218
|
229 |
|
haftmann@24589
|
230 |
fun gen_instance prep_arity read_def do_proof raw_arities raw_defs after_qed theory =
|
haftmann@24218
|
231 |
let
|
haftmann@24218
|
232 |
val arities = map (prep_arity theory) raw_arities;
|
wenzelm@25020
|
233 |
val _ = if null arities then error "At least one arity must be given" else ();
|
haftmann@24218
|
234 |
val _ = case (duplicates (op =) o map #1) arities
|
haftmann@24218
|
235 |
of [] => ()
|
wenzelm@25020
|
236 |
| dupl_tycos => error ("Type constructors occur more than once in arities: "
|
wenzelm@25020
|
237 |
^ commas_quote dupl_tycos);
|
haftmann@24218
|
238 |
fun get_consts_class tyco ty class =
|
haftmann@24218
|
239 |
let
|
wenzelm@24968
|
240 |
val cs = (these o try (#params o AxClass.get_info theory)) class;
|
haftmann@24218
|
241 |
val subst_ty = map_type_tfree (K ty);
|
haftmann@24218
|
242 |
in
|
haftmann@24304
|
243 |
map (fn (c, ty) => (c, ((class, tyco), subst_ty ty))) cs
|
haftmann@24218
|
244 |
end;
|
haftmann@24218
|
245 |
fun get_consts_sort (tyco, asorts, sort) =
|
haftmann@24218
|
246 |
let
|
haftmann@24589
|
247 |
val ty = Type (tyco, map (fn (v, sort) => TVar ((v, 0), sort))
|
wenzelm@24847
|
248 |
(Name.names Name.context Name.aT asorts))
|
wenzelm@24731
|
249 |
in maps (get_consts_class tyco ty) (Sign.complete_sort theory sort) end;
|
haftmann@24218
|
250 |
val cs = maps get_consts_sort arities;
|
haftmann@24218
|
251 |
fun mk_typnorm thy (ty, ty_sc) =
|
haftmann@24218
|
252 |
case try (Sign.typ_match thy (Logic.varifyT ty_sc, ty)) Vartab.empty
|
haftmann@24218
|
253 |
of SOME env => SOME (Logic.varifyT #> Envir.typ_subst_TVars env #> Logic.unvarifyT)
|
haftmann@24218
|
254 |
| NONE => NONE;
|
haftmann@24218
|
255 |
fun read_defs defs cs thy_read =
|
haftmann@24218
|
256 |
let
|
haftmann@24218
|
257 |
fun read raw_def cs =
|
haftmann@24218
|
258 |
let
|
haftmann@24218
|
259 |
val (c, (inst, ((name_opt, t), atts))) = read_def thy_read raw_def;
|
haftmann@24218
|
260 |
val ty = Consts.instance (Sign.consts_of thy_read) (c, inst);
|
haftmann@24304
|
261 |
val ((class, tyco), ty') = case AList.lookup (op =) cs c
|
wenzelm@25020
|
262 |
of NONE => error ("Illegal definition for constant " ^ quote c)
|
haftmann@24218
|
263 |
| SOME class_ty => class_ty;
|
haftmann@24218
|
264 |
val name = case name_opt
|
haftmann@24218
|
265 |
of NONE => Thm.def_name (Logic.name_arity (tyco, [], c))
|
haftmann@24218
|
266 |
| SOME name => name;
|
haftmann@24218
|
267 |
val t' = case mk_typnorm thy_read (ty', ty)
|
wenzelm@25020
|
268 |
of NONE => error ("Illegal definition for constant " ^
|
haftmann@24218
|
269 |
quote (c ^ "::" ^ setmp show_sorts true
|
haftmann@24218
|
270 |
(Sign.string_of_typ thy_read) ty))
|
haftmann@24218
|
271 |
| SOME norm => map_types norm t
|
haftmann@24218
|
272 |
in (((class, tyco), ((name, t'), atts)), AList.delete (op =) c cs) end;
|
haftmann@24218
|
273 |
in fold_map read defs cs end;
|
haftmann@24304
|
274 |
val (defs, other_cs) = read_defs raw_defs cs
|
haftmann@24218
|
275 |
(fold Sign.primitive_arity arities (Theory.copy theory));
|
haftmann@24423
|
276 |
fun after_qed' cs defs =
|
wenzelm@24766
|
277 |
fold Sign.add_const_constraint (map (apsnd SOME) cs)
|
haftmann@24423
|
278 |
#> after_qed defs;
|
haftmann@24218
|
279 |
in
|
haftmann@24218
|
280 |
theory
|
haftmann@25038
|
281 |
|> fold_map get_remove_global_constraint (map fst cs |> distinct (op =))
|
haftmann@24304
|
282 |
||>> fold_map add_def defs
|
haftmann@24304
|
283 |
||> fold (fn (c, ((class, tyco), ty)) => add_inst_def' (class, tyco) (c, ty)) other_cs
|
haftmann@24423
|
284 |
|-> (fn (cs, defs) => do_proof (after_qed' cs defs) arities defs)
|
haftmann@24218
|
285 |
end;
|
haftmann@24218
|
286 |
|
haftmann@24423
|
287 |
fun tactic_proof tac f arities defs =
|
haftmann@24218
|
288 |
fold (fn arity => AxClass.prove_arity arity tac) arities
|
haftmann@24423
|
289 |
#> f
|
haftmann@24423
|
290 |
#> pair defs;
|
haftmann@24218
|
291 |
|
haftmann@24218
|
292 |
in
|
haftmann@24218
|
293 |
|
haftmann@24589
|
294 |
val instance = gen_instance Sign.cert_arity read_def
|
haftmann@24589
|
295 |
(fn f => fn arities => fn defs => instance_arity f arities);
|
haftmann@24589
|
296 |
val instance_cmd = gen_instance Sign.read_arity read_def_cmd
|
haftmann@24589
|
297 |
(fn f => fn arities => fn defs => instance_arity f arities);
|
haftmann@24589
|
298 |
fun prove_instance tac arities defs =
|
haftmann@24589
|
299 |
gen_instance Sign.cert_arity read_def
|
haftmann@24589
|
300 |
(tactic_proof tac) arities defs (K I);
|
haftmann@24218
|
301 |
|
haftmann@24218
|
302 |
end; (*local*)
|
haftmann@24218
|
303 |
|
haftmann@24218
|
304 |
|
haftmann@24218
|
305 |
|
haftmann@24589
|
306 |
(** class data **)
|
haftmann@24218
|
307 |
|
haftmann@24218
|
308 |
datatype class_data = ClassData of {
|
haftmann@24218
|
309 |
consts: (string * string) list
|
haftmann@24836
|
310 |
(*locale parameter ~> constant name*),
|
haftmann@25062
|
311 |
base_sort: sort,
|
haftmann@25083
|
312 |
inst: term option list
|
haftmann@25083
|
313 |
(*canonical interpretation*),
|
haftmann@25062
|
314 |
morphism: morphism,
|
haftmann@25062
|
315 |
(*partial morphism of canonical interpretation*)
|
haftmann@24657
|
316 |
intro: thm,
|
haftmann@24657
|
317 |
defs: thm list,
|
haftmann@25062
|
318 |
operations: (string * ((term * typ) * (typ * int))) list
|
haftmann@25062
|
319 |
(*constant name ~> (locale term & typ,
|
haftmann@25062
|
320 |
(constant constraint, instantiaton index of class typ))*)
|
haftmann@24657
|
321 |
};
|
haftmann@24218
|
322 |
|
haftmann@24657
|
323 |
fun rep_class_data (ClassData d) = d;
|
haftmann@25062
|
324 |
fun mk_class_data ((consts, base_sort, inst, morphism, intro),
|
haftmann@25062
|
325 |
(defs, operations)) =
|
haftmann@25062
|
326 |
ClassData { consts = consts, base_sort = base_sort, inst = inst,
|
haftmann@25062
|
327 |
morphism = morphism, intro = intro, defs = defs,
|
haftmann@25062
|
328 |
operations = operations };
|
haftmann@25062
|
329 |
fun map_class_data f (ClassData { consts, base_sort, inst, morphism, intro,
|
haftmann@25062
|
330 |
defs, operations }) =
|
haftmann@25062
|
331 |
mk_class_data (f ((consts, base_sort, inst, morphism, intro),
|
haftmann@25062
|
332 |
(defs, operations)));
|
haftmann@25038
|
333 |
fun merge_class_data _ (ClassData { consts = consts,
|
haftmann@25062
|
334 |
base_sort = base_sort, inst = inst, morphism = morphism, intro = intro,
|
haftmann@25062
|
335 |
defs = defs1, operations = operations1 },
|
haftmann@25062
|
336 |
ClassData { consts = _, base_sort = _, inst = _, morphism = _, intro = _,
|
haftmann@25062
|
337 |
defs = defs2, operations = operations2 }) =
|
haftmann@25062
|
338 |
mk_class_data ((consts, base_sort, inst, morphism, intro),
|
haftmann@24914
|
339 |
(Thm.merge_thms (defs1, defs2),
|
haftmann@25062
|
340 |
AList.merge (op =) (K true) (operations1, operations2)));
|
haftmann@24218
|
341 |
|
haftmann@24218
|
342 |
structure ClassData = TheoryDataFun
|
haftmann@24218
|
343 |
(
|
haftmann@25038
|
344 |
type T = class_data Graph.T
|
haftmann@25038
|
345 |
val empty = Graph.empty;
|
haftmann@24218
|
346 |
val copy = I;
|
haftmann@24218
|
347 |
val extend = I;
|
haftmann@25038
|
348 |
fun merge _ = Graph.join merge_class_data;
|
haftmann@24218
|
349 |
);
|
haftmann@24218
|
350 |
|
haftmann@24218
|
351 |
|
haftmann@24218
|
352 |
(* queries *)
|
haftmann@24218
|
353 |
|
haftmann@25038
|
354 |
val lookup_class_data = Option.map rep_class_data oo try o Graph.get_node o ClassData.get;
|
haftmann@24218
|
355 |
|
haftmann@24589
|
356 |
fun the_class_data thy class = case lookup_class_data thy class
|
wenzelm@25020
|
357 |
of NONE => error ("Undeclared class " ^ quote class)
|
haftmann@24589
|
358 |
| SOME data => data;
|
haftmann@24218
|
359 |
|
haftmann@25038
|
360 |
val is_class = is_some oo lookup_class_data;
|
haftmann@25038
|
361 |
|
haftmann@25038
|
362 |
val ancestry = Graph.all_succs o ClassData.get;
|
haftmann@24218
|
363 |
|
haftmann@25002
|
364 |
fun these_params thy =
|
haftmann@24218
|
365 |
let
|
haftmann@24218
|
366 |
fun params class =
|
haftmann@24218
|
367 |
let
|
wenzelm@24930
|
368 |
val const_typs = (#params o AxClass.get_info thy) class;
|
haftmann@24657
|
369 |
val const_names = (#consts o the_class_data thy) class;
|
haftmann@24218
|
370 |
in
|
haftmann@24218
|
371 |
(map o apsnd) (fn c => (c, (the o AList.lookup (op =) const_typs) c)) const_names
|
haftmann@24218
|
372 |
end;
|
haftmann@24218
|
373 |
in maps params o ancestry thy end;
|
haftmann@24218
|
374 |
|
haftmann@24657
|
375 |
fun these_defs thy = maps (these o Option.map #defs o lookup_class_data thy) o ancestry thy;
|
haftmann@24218
|
376 |
|
haftmann@25062
|
377 |
fun morphism thy = #morphism o the_class_data thy;
|
haftmann@25062
|
378 |
|
haftmann@24218
|
379 |
fun these_intros thy =
|
haftmann@24657
|
380 |
Graph.fold (fn (_, (data, _)) => insert Thm.eq_thm ((#intro o rep_class_data) data))
|
haftmann@25038
|
381 |
(ClassData.get thy) [];
|
haftmann@24218
|
382 |
|
haftmann@24836
|
383 |
fun these_operations thy =
|
haftmann@24836
|
384 |
maps (#operations o the_class_data thy) o ancestry thy;
|
haftmann@24657
|
385 |
|
haftmann@25002
|
386 |
fun local_operation thy = AList.lookup (op =) o these_operations thy;
|
haftmann@25002
|
387 |
|
haftmann@25062
|
388 |
fun sups_base_sort thy sort =
|
haftmann@24901
|
389 |
let
|
haftmann@25062
|
390 |
val sups = filter (is_class thy) sort
|
haftmann@24901
|
391 |
|> Sign.minimize_sort thy;
|
haftmann@25062
|
392 |
val base_sort = case sups
|
haftmann@25062
|
393 |
of _ :: _ => maps (#base_sort o the_class_data thy) sups
|
haftmann@25062
|
394 |
|> Sign.minimize_sort thy
|
haftmann@24901
|
395 |
| [] => sort;
|
haftmann@25062
|
396 |
in (sups, base_sort) end;
|
haftmann@24901
|
397 |
|
haftmann@24218
|
398 |
fun print_classes thy =
|
haftmann@24218
|
399 |
let
|
wenzelm@24920
|
400 |
val ctxt = ProofContext.init thy;
|
haftmann@24218
|
401 |
val algebra = Sign.classes_of thy;
|
haftmann@24218
|
402 |
val arities =
|
haftmann@24218
|
403 |
Symtab.empty
|
haftmann@24218
|
404 |
|> Symtab.fold (fn (tyco, arities) => fold (fn (class, _) =>
|
haftmann@24218
|
405 |
Symtab.map_default (class, []) (insert (op =) tyco)) arities)
|
haftmann@24218
|
406 |
((#arities o Sorts.rep_algebra) algebra);
|
haftmann@24218
|
407 |
val the_arities = these o Symtab.lookup arities;
|
haftmann@24218
|
408 |
fun mk_arity class tyco =
|
haftmann@24218
|
409 |
let
|
haftmann@24218
|
410 |
val Ss = Sorts.mg_domain algebra tyco [class];
|
wenzelm@24920
|
411 |
in Syntax.pretty_arity ctxt (tyco, Ss, [class]) end;
|
haftmann@24218
|
412 |
fun mk_param (c, ty) = Pretty.str (Sign.extern_const thy c ^ " :: "
|
wenzelm@24920
|
413 |
^ setmp show_sorts false (Syntax.string_of_typ ctxt o Type.strip_sorts) ty);
|
haftmann@24218
|
414 |
fun mk_entry class = (Pretty.block o Pretty.fbreaks o map_filter I) [
|
haftmann@25062
|
415 |
(SOME o Pretty.str) ("class " ^ Sign.extern_class thy class ^ ":"),
|
haftmann@24218
|
416 |
(SOME o Pretty.block) [Pretty.str "supersort: ",
|
wenzelm@24920
|
417 |
(Syntax.pretty_sort ctxt o Sign.minimize_sort thy o Sign.super_classes thy) class],
|
haftmann@25062
|
418 |
if is_class thy class then (SOME o Pretty.str)
|
haftmann@25062
|
419 |
("locale: " ^ Locale.extern thy class) else NONE,
|
haftmann@25062
|
420 |
((fn [] => NONE | ps => (SOME o Pretty.block o Pretty.fbreaks)
|
haftmann@25062
|
421 |
(Pretty.str "parameters:" :: ps)) o map mk_param
|
wenzelm@24930
|
422 |
o these o Option.map #params o try (AxClass.get_info thy)) class,
|
haftmann@24218
|
423 |
(SOME o Pretty.block o Pretty.breaks) [
|
haftmann@24218
|
424 |
Pretty.str "instances:",
|
haftmann@24218
|
425 |
Pretty.list "" "" (map (mk_arity class) (the_arities class))
|
haftmann@24218
|
426 |
]
|
haftmann@24218
|
427 |
]
|
haftmann@24218
|
428 |
in
|
haftmann@24589
|
429 |
(Pretty.writeln o Pretty.chunks o separate (Pretty.str "")
|
haftmann@24589
|
430 |
o map mk_entry o Sorts.all_classes) algebra
|
haftmann@24218
|
431 |
end;
|
haftmann@24218
|
432 |
|
haftmann@24218
|
433 |
|
haftmann@24218
|
434 |
(* updaters *)
|
haftmann@24218
|
435 |
|
haftmann@25083
|
436 |
fun add_class_data ((class, superclasses), (cs, base_sort, insttab, phi, intro)) thy =
|
haftmann@25002
|
437 |
let
|
haftmann@25083
|
438 |
val inst = map
|
haftmann@25083
|
439 |
(SOME o the o Symtab.lookup insttab o fst o fst)
|
haftmann@25038
|
440 |
(Locale.parameters_of_expr thy (Locale.Locale class));
|
haftmann@25062
|
441 |
val operations = map (fn (v_ty as (_, ty'), (c, ty)) =>
|
haftmann@25062
|
442 |
(c, ((Free v_ty, ty'), (Logic.varifyT ty, 0)))) cs;
|
haftmann@25002
|
443 |
val cs = (map o pairself) fst cs;
|
haftmann@25038
|
444 |
val add_class = Graph.new_node (class,
|
haftmann@25083
|
445 |
mk_class_data ((cs, base_sort, inst, phi, intro), ([], operations)))
|
haftmann@25002
|
446 |
#> fold (curry Graph.add_edge class) superclasses;
|
haftmann@25002
|
447 |
in
|
haftmann@25038
|
448 |
ClassData.map add_class thy
|
haftmann@25002
|
449 |
end;
|
haftmann@24218
|
450 |
|
haftmann@25062
|
451 |
fun register_operation class ((c, rhs), some_def) thy =
|
haftmann@25062
|
452 |
let
|
haftmann@25062
|
453 |
val ty = Sign.the_const_constraint thy c;
|
haftmann@25062
|
454 |
val typargs = Sign.const_typargs thy (c, ty);
|
haftmann@25062
|
455 |
val typidx = find_index (fn TVar ((v, _), _) => Name.aT = v | _ => false) typargs;
|
haftmann@25062
|
456 |
val ty' = Term.fastype_of rhs;
|
haftmann@25062
|
457 |
in
|
haftmann@25062
|
458 |
thy
|
haftmann@25062
|
459 |
|> (ClassData.map o Graph.map_node class o map_class_data o apsnd)
|
haftmann@25062
|
460 |
(fn (defs, operations) =>
|
haftmann@25062
|
461 |
(case some_def of NONE => defs | SOME def => def :: defs,
|
haftmann@25062
|
462 |
(c, ((rhs, ty'), (ty, typidx))) :: operations))
|
haftmann@25062
|
463 |
end;
|
haftmann@24218
|
464 |
|
haftmann@24589
|
465 |
|
haftmann@24589
|
466 |
(** rule calculation, tactics and methods **)
|
haftmann@24589
|
467 |
|
wenzelm@25024
|
468 |
val class_prefix = Logic.const_of_class o Sign.base_name;
|
wenzelm@25024
|
469 |
|
haftmann@25062
|
470 |
fun calculate_morphism class cs =
|
haftmann@25062
|
471 |
let
|
haftmann@25062
|
472 |
val subst_typ = Term.map_type_tfree (fn var as (v, sort) =>
|
haftmann@25062
|
473 |
if v = Name.aT then TVar ((v, 0), [class]) else TVar ((v, 0), sort));
|
haftmann@25062
|
474 |
fun subst_aterm (t as Free (v, ty)) = (case AList.lookup (op =) cs v
|
haftmann@25062
|
475 |
of SOME (c, _) => Const (c, ty)
|
haftmann@25062
|
476 |
| NONE => t)
|
haftmann@25062
|
477 |
| subst_aterm t = t;
|
haftmann@25062
|
478 |
val subst_term = map_aterms subst_aterm #> map_types subst_typ;
|
haftmann@25062
|
479 |
in
|
haftmann@25062
|
480 |
Morphism.identity
|
haftmann@25062
|
481 |
$> Morphism.term_morphism subst_term
|
haftmann@25062
|
482 |
$> Morphism.typ_morphism subst_typ
|
haftmann@25062
|
483 |
end;
|
haftmann@25062
|
484 |
|
haftmann@25038
|
485 |
fun class_intro thy class sups =
|
haftmann@24589
|
486 |
let
|
haftmann@24589
|
487 |
fun class_elim class =
|
wenzelm@25020
|
488 |
case (#axioms o AxClass.get_info thy) class
|
wenzelm@25020
|
489 |
of [thm] => SOME (Drule.unconstrainTs thm)
|
haftmann@24589
|
490 |
| [] => NONE;
|
haftmann@25038
|
491 |
val pred_intro = case Locale.intros thy class
|
haftmann@24589
|
492 |
of ([ax_intro], [intro]) => intro |> OF_LAST ax_intro |> SOME
|
haftmann@24589
|
493 |
| ([intro], []) => SOME intro
|
haftmann@24589
|
494 |
| ([], [intro]) => SOME intro
|
haftmann@24589
|
495 |
| _ => NONE;
|
haftmann@24589
|
496 |
val pred_intro' = pred_intro
|
haftmann@24589
|
497 |
|> Option.map (fn intro => intro OF map_filter class_elim sups);
|
wenzelm@24930
|
498 |
val class_intro = (#intro o AxClass.get_info thy) class;
|
haftmann@24589
|
499 |
val raw_intro = case pred_intro'
|
haftmann@24589
|
500 |
of SOME pred_intro => class_intro |> OF_LAST pred_intro
|
haftmann@24589
|
501 |
| NONE => class_intro;
|
haftmann@24589
|
502 |
val sort = Sign.super_classes thy class;
|
wenzelm@24847
|
503 |
val typ = TVar ((Name.aT, 0), sort);
|
haftmann@24589
|
504 |
val defs = these_defs thy sups;
|
haftmann@24589
|
505 |
in
|
haftmann@24589
|
506 |
raw_intro
|
haftmann@24589
|
507 |
|> Drule.instantiate' [SOME (Thm.ctyp_of thy typ)] []
|
haftmann@24589
|
508 |
|> strip_all_ofclass thy sort
|
haftmann@24589
|
509 |
|> Thm.strip_shyps
|
haftmann@24589
|
510 |
|> MetaSimplifier.rewrite_rule defs
|
haftmann@24589
|
511 |
|> Drule.unconstrainTs
|
haftmann@24589
|
512 |
end;
|
haftmann@24589
|
513 |
|
haftmann@24589
|
514 |
fun class_interpretation class facts defs thy =
|
haftmann@24589
|
515 |
let
|
haftmann@25038
|
516 |
val params = these_params thy [class];
|
haftmann@25083
|
517 |
val inst = (#inst o the_class_data thy) class;
|
wenzelm@25020
|
518 |
val tac = ALLGOALS (ProofContext.fact_tac facts);
|
haftmann@25038
|
519 |
val prfx = class_prefix class;
|
haftmann@24589
|
520 |
in
|
haftmann@25038
|
521 |
thy
|
haftmann@25038
|
522 |
|> fold_map (get_remove_global_constraint o fst o snd) params
|
ballarin@25094
|
523 |
||> prove_interpretation tac ((false, prfx), []) (Locale.Locale class)
|
ballarin@25094
|
524 |
(inst, map (fn def => (("", []), def)) defs)
|
haftmann@25038
|
525 |
|-> (fn cs => fold (Sign.add_const_constraint o apsnd SOME) cs)
|
haftmann@24589
|
526 |
end;
|
haftmann@24218
|
527 |
|
haftmann@24218
|
528 |
fun intro_classes_tac facts st =
|
haftmann@24218
|
529 |
let
|
haftmann@24218
|
530 |
val thy = Thm.theory_of_thm st;
|
haftmann@24218
|
531 |
val classes = Sign.all_classes thy;
|
haftmann@24218
|
532 |
val class_trivs = map (Thm.class_triv thy) classes;
|
haftmann@24218
|
533 |
val class_intros = these_intros thy;
|
wenzelm@24930
|
534 |
val axclass_intros = map_filter (try (#intro o AxClass.get_info thy)) classes;
|
haftmann@24218
|
535 |
in
|
wenzelm@24930
|
536 |
(ALLGOALS (Method.insert_tac facts THEN'
|
wenzelm@24930
|
537 |
REPEAT_ALL_NEW (resolve_tac (class_trivs @ class_intros @ axclass_intros)))
|
wenzelm@24930
|
538 |
THEN Tactic.distinct_subgoals_tac) st
|
haftmann@24218
|
539 |
end;
|
haftmann@24218
|
540 |
|
haftmann@24218
|
541 |
fun default_intro_classes_tac [] = intro_classes_tac []
|
wenzelm@24930
|
542 |
| default_intro_classes_tac _ = no_tac;
|
haftmann@24218
|
543 |
|
haftmann@24218
|
544 |
fun default_tac rules ctxt facts =
|
haftmann@24218
|
545 |
HEADGOAL (Method.some_rule_tac rules ctxt facts) ORELSE
|
haftmann@24218
|
546 |
default_intro_classes_tac facts;
|
haftmann@24218
|
547 |
|
haftmann@24218
|
548 |
val _ = Context.add_setup (Method.add_methods
|
haftmann@24218
|
549 |
[("intro_classes", Method.no_args (Method.METHOD intro_classes_tac),
|
haftmann@24218
|
550 |
"back-chain introduction rules of classes"),
|
haftmann@24218
|
551 |
("default", Method.thms_ctxt_args (Method.METHOD oo default_tac),
|
haftmann@24218
|
552 |
"apply some intro/elim rule")]);
|
haftmann@24218
|
553 |
|
haftmann@24218
|
554 |
|
haftmann@24589
|
555 |
(** classes and class target **)
|
haftmann@24218
|
556 |
|
haftmann@25002
|
557 |
(* class context syntax *)
|
haftmann@24748
|
558 |
|
haftmann@25083
|
559 |
structure ClassSyntax = ProofDataFun(
|
haftmann@25083
|
560 |
type T = {
|
haftmann@25083
|
561 |
constraints: (string * typ) list,
|
haftmann@25083
|
562 |
base_sort: sort,
|
haftmann@25083
|
563 |
local_operation: string * typ -> (typ * term) option,
|
haftmann@25083
|
564 |
rews: (term * term) list,
|
haftmann@25083
|
565 |
passed: bool
|
haftmann@25083
|
566 |
} option;
|
haftmann@25083
|
567 |
fun init _ = NONE;
|
haftmann@25083
|
568 |
);
|
haftmann@25083
|
569 |
|
haftmann@25083
|
570 |
fun synchronize_syntax thy sups base_sort ctxt =
|
haftmann@24914
|
571 |
let
|
haftmann@25083
|
572 |
val operations = these_operations thy sups;
|
haftmann@25083
|
573 |
|
haftmann@25083
|
574 |
(* constraints *)
|
haftmann@25083
|
575 |
fun local_constraint (c, (_, (ty, _))) =
|
haftmann@25083
|
576 |
let
|
haftmann@25083
|
577 |
val ty' = ty
|
haftmann@25083
|
578 |
|> map_atyps (fn ty as TVar ((v, 0), _) =>
|
haftmann@25083
|
579 |
if v = Name.aT then TVar ((v, 0), base_sort) else ty)
|
haftmann@25083
|
580 |
|> SOME;
|
haftmann@25083
|
581 |
in (c, ty') end
|
haftmann@25083
|
582 |
val constraints = (map o apsnd) (fst o snd) operations;
|
haftmann@25083
|
583 |
|
haftmann@25083
|
584 |
(* check phase *)
|
haftmann@25083
|
585 |
val typargs = Consts.typargs (ProofContext.consts_of ctxt);
|
haftmann@25083
|
586 |
fun check_const (c, ty) ((t, _), (_, idx)) =
|
haftmann@25083
|
587 |
((nth (typargs (c, ty)) idx), t);
|
haftmann@25083
|
588 |
fun local_operation (c_ty as (c, _)) = AList.lookup (op =) operations c
|
haftmann@25083
|
589 |
|> Option.map (check_const c_ty);
|
haftmann@25002
|
590 |
|
haftmann@25083
|
591 |
(* uncheck phase *)
|
haftmann@25083
|
592 |
val proto_rews = map (fn (c, ((t, ty), _)) => (t, Const (c, ty))) operations;
|
haftmann@25083
|
593 |
fun rew_app f (t1 $ t2) = rew_app f t1 $ f t2
|
haftmann@25083
|
594 |
| rew_app f t = t;
|
haftmann@25083
|
595 |
val rews = (map o apfst o rew_app)
|
haftmann@25083
|
596 |
(Pattern.rewrite_term thy proto_rews []) proto_rews;
|
haftmann@25083
|
597 |
in
|
haftmann@25083
|
598 |
ctxt
|
haftmann@25083
|
599 |
|> fold (ProofContext.add_const_constraint o local_constraint) operations
|
haftmann@25083
|
600 |
|> ClassSyntax.map (K (SOME {
|
haftmann@25083
|
601 |
constraints = constraints,
|
haftmann@25083
|
602 |
base_sort = base_sort,
|
haftmann@25083
|
603 |
local_operation = local_operation,
|
haftmann@25083
|
604 |
rews = rews,
|
haftmann@25083
|
605 |
passed = false
|
haftmann@25083
|
606 |
}))
|
haftmann@25083
|
607 |
end;
|
haftmann@25083
|
608 |
|
haftmann@25083
|
609 |
fun refresh_syntax class ctxt =
|
haftmann@25002
|
610 |
let
|
haftmann@25002
|
611 |
val thy = ProofContext.theory_of ctxt;
|
haftmann@25062
|
612 |
val base_sort = (#base_sort o the_class_data thy) class;
|
haftmann@25083
|
613 |
in synchronize_syntax thy [class] base_sort ctxt end;
|
haftmann@24914
|
614 |
|
haftmann@25083
|
615 |
val mark_passed = (ClassSyntax.map o Option.map)
|
haftmann@25083
|
616 |
(fn { constraints, base_sort, local_operation, rews, passed } =>
|
haftmann@25083
|
617 |
{ constraints = constraints, base_sort = base_sort,
|
haftmann@25083
|
618 |
local_operation = local_operation, rews = rews, passed = true });
|
haftmann@25083
|
619 |
|
haftmann@25083
|
620 |
fun sort_term_check ts ctxt =
|
haftmann@24748
|
621 |
let
|
haftmann@25083
|
622 |
val { constraints, base_sort, local_operation, passed, ... } =
|
haftmann@25083
|
623 |
the (ClassSyntax.get ctxt);
|
haftmann@25083
|
624 |
fun check_typ (c, ty) (TFree (v, _), t) = if v = Name.aT
|
haftmann@25062
|
625 |
then apfst (AList.update (op =) ((c, ty), t)) else I
|
haftmann@25083
|
626 |
| check_typ (c, ty) (TVar (vi, _), t) = if TypeInfer.is_param vi
|
haftmann@25062
|
627 |
then apfst (AList.update (op =) ((c, ty), t))
|
haftmann@24901
|
628 |
#> apsnd (insert (op =) vi) else I
|
haftmann@25083
|
629 |
| check_typ _ _ = I;
|
haftmann@25083
|
630 |
fun add_const (Const c_ty) = Option.map (check_typ c_ty) (local_operation c_ty)
|
haftmann@25062
|
631 |
|> the_default I
|
haftmann@24901
|
632 |
| add_const _ = I;
|
haftmann@24901
|
633 |
val (cs, typarams) = (fold o fold_aterms) add_const ts ([], []);
|
haftmann@25062
|
634 |
val subst_typ = map_type_tvar (fn var as (vi, _) =>
|
haftmann@25062
|
635 |
if member (op =) typarams vi then TFree (Name.aT, base_sort) else TVar var);
|
haftmann@25062
|
636 |
val subst_term = map_aterms
|
haftmann@24901
|
637 |
(fn t as Const (c, ty) => the_default t (AList.lookup (op =) cs (c, ty)) | t => t)
|
haftmann@25062
|
638 |
#> map_types subst_typ;
|
haftmann@25062
|
639 |
val ts' = map subst_term ts;
|
haftmann@25083
|
640 |
in if eq_list (op aconv) (ts, ts') andalso passed then NONE
|
haftmann@25083
|
641 |
else
|
haftmann@25083
|
642 |
ctxt
|
haftmann@25083
|
643 |
|> fold (ProofContext.add_const_constraint o apsnd SOME) constraints
|
haftmann@25083
|
644 |
|> mark_passed
|
haftmann@25083
|
645 |
|> pair ts'
|
haftmann@25083
|
646 |
|> SOME
|
haftmann@25083
|
647 |
end;
|
haftmann@24748
|
648 |
|
haftmann@25062
|
649 |
val uncheck = ref true;
|
haftmann@25002
|
650 |
|
haftmann@25083
|
651 |
fun sort_term_uncheck ts ctxt =
|
haftmann@25002
|
652 |
let
|
haftmann@25062
|
653 |
(*FIXME abbreviations*)
|
haftmann@25002
|
654 |
val thy = ProofContext.theory_of ctxt;
|
haftmann@25083
|
655 |
val rews = (#rews o the o ClassSyntax.get) ctxt;
|
haftmann@25062
|
656 |
val ts' = if ! uncheck
|
haftmann@25083
|
657 |
then map (Pattern.rewrite_term thy rews []) ts else ts;
|
wenzelm@25060
|
658 |
in if eq_list (op aconv) (ts, ts') then NONE else SOME (ts', ctxt) end;
|
haftmann@25002
|
659 |
|
haftmann@25083
|
660 |
fun init_ctxt thy sups base_sort ctxt =
|
haftmann@25083
|
661 |
ctxt
|
haftmann@25083
|
662 |
|> Variable.declare_term
|
haftmann@25083
|
663 |
(Logic.mk_type (TFree (Name.aT, base_sort)))
|
haftmann@25083
|
664 |
|> synchronize_syntax thy sups base_sort
|
haftmann@25083
|
665 |
|> Context.proof_map (
|
haftmann@25083
|
666 |
Syntax.add_term_check 0 "class" sort_term_check
|
haftmann@25083
|
667 |
#> Syntax.add_term_uncheck (~10) "class" sort_term_uncheck)
|
haftmann@24901
|
668 |
|
haftmann@24901
|
669 |
fun init class ctxt =
|
haftmann@25083
|
670 |
let
|
haftmann@25083
|
671 |
val thy = ProofContext.theory_of ctxt;
|
haftmann@25083
|
672 |
in
|
haftmann@25083
|
673 |
init_ctxt thy [class] ((#base_sort o the_class_data thy) class) ctxt
|
haftmann@25083
|
674 |
end;
|
haftmann@24914
|
675 |
|
haftmann@24748
|
676 |
|
haftmann@24589
|
677 |
(* class definition *)
|
haftmann@24218
|
678 |
|
haftmann@24218
|
679 |
local
|
haftmann@24218
|
680 |
|
haftmann@24748
|
681 |
fun gen_class_spec prep_class prep_expr process_expr thy raw_supclasses raw_includes_elems =
|
haftmann@24218
|
682 |
let
|
haftmann@24748
|
683 |
val supclasses = map (prep_class thy) raw_supclasses;
|
haftmann@25062
|
684 |
val (sups, base_sort) = sups_base_sort thy supclasses;
|
haftmann@24748
|
685 |
val supsort = Sign.minimize_sort thy supclasses;
|
haftmann@25038
|
686 |
val suplocales = map Locale.Locale sups;
|
haftmann@24748
|
687 |
val (raw_elems, includes) = fold_rev (fn Locale.Elem e => apfst (cons e)
|
haftmann@24748
|
688 |
| Locale.Expr i => apsnd (cons (prep_expr thy i))) raw_includes_elems ([], []);
|
haftmann@24748
|
689 |
val supexpr = Locale.Merge suplocales;
|
haftmann@24748
|
690 |
val supparams = (map fst o Locale.parameters_of_expr thy) supexpr;
|
haftmann@25002
|
691 |
val supconsts = AList.make (the o AList.lookup (op =) (these_params thy sups))
|
haftmann@24748
|
692 |
(map fst supparams);
|
haftmann@24748
|
693 |
val mergeexpr = Locale.Merge (suplocales @ includes);
|
haftmann@24748
|
694 |
val constrain = Element.Constrains ((map o apsnd o map_atyps)
|
wenzelm@24847
|
695 |
(fn TFree (_, sort) => TFree (Name.aT, sort)) supparams);
|
haftmann@24748
|
696 |
in
|
haftmann@24748
|
697 |
ProofContext.init thy
|
haftmann@24748
|
698 |
|> Locale.cert_expr supexpr [constrain]
|
haftmann@24748
|
699 |
|> snd
|
haftmann@25083
|
700 |
|> init_ctxt thy sups base_sort
|
haftmann@24748
|
701 |
|> process_expr Locale.empty raw_elems
|
haftmann@24748
|
702 |
|> fst
|
haftmann@25062
|
703 |
|> (fn elems => ((((sups, supconsts), (supsort, base_sort, mergeexpr)),
|
haftmann@24748
|
704 |
(*FIXME*) if null includes then constrain :: elems else elems)))
|
haftmann@24748
|
705 |
end;
|
haftmann@24748
|
706 |
|
haftmann@24748
|
707 |
val read_class_spec = gen_class_spec Sign.intern_class Locale.intern_expr Locale.read_expr;
|
haftmann@24748
|
708 |
val check_class_spec = gen_class_spec (K I) (K I) Locale.cert_expr;
|
haftmann@24748
|
709 |
|
wenzelm@24968
|
710 |
fun define_class_params (name, raw_superclasses) raw_consts raw_dep_axioms other_consts thy =
|
wenzelm@24968
|
711 |
let
|
wenzelm@24968
|
712 |
val superclasses = map (Sign.certify_class thy) raw_superclasses;
|
wenzelm@24968
|
713 |
val consts = (map o apfst o apsnd) (Sign.certify_typ thy) raw_consts;
|
haftmann@25083
|
714 |
fun add_const ((c, ty), syn) =
|
haftmann@25083
|
715 |
Sign.declare_const [] (c, Type.strip_sorts ty, syn) #>> Term.dest_Const;
|
wenzelm@24968
|
716 |
fun mk_axioms cs thy =
|
wenzelm@24968
|
717 |
raw_dep_axioms thy cs
|
wenzelm@24968
|
718 |
|> (map o apsnd o map) (Sign.cert_prop thy)
|
wenzelm@24968
|
719 |
|> rpair thy;
|
haftmann@25002
|
720 |
fun constrain_typs class = (map o apsnd o Term.map_type_tfree)
|
haftmann@25002
|
721 |
(fn (v, _) => TFree (v, [class]))
|
wenzelm@24968
|
722 |
in
|
wenzelm@24968
|
723 |
thy
|
wenzelm@24968
|
724 |
|> Sign.add_path (Logic.const_of_class name)
|
wenzelm@24968
|
725 |
|> fold_map add_const consts
|
wenzelm@24968
|
726 |
||> Sign.restore_naming thy
|
wenzelm@24968
|
727 |
|-> (fn cs => mk_axioms cs
|
wenzelm@24968
|
728 |
#-> (fn axioms_prop => AxClass.define_class (name, superclasses)
|
wenzelm@24968
|
729 |
(map fst cs @ other_consts) axioms_prop
|
haftmann@25002
|
730 |
#-> (fn class => `(fn _ => constrain_typs class cs)
|
haftmann@25002
|
731 |
#-> (fn cs' => `(fn thy => AxClass.get_info thy class)
|
haftmann@25002
|
732 |
#-> (fn {axioms, ...} => fold (Sign.add_const_constraint o apsnd SOME) cs'
|
haftmann@25002
|
733 |
#> pair (class, (cs', axioms)))))))
|
wenzelm@24968
|
734 |
end;
|
wenzelm@24968
|
735 |
|
haftmann@25002
|
736 |
fun gen_class prep_spec prep_param bname
|
haftmann@24748
|
737 |
raw_supclasses raw_includes_elems raw_other_consts thy =
|
haftmann@24748
|
738 |
let
|
haftmann@25038
|
739 |
val class = Sign.full_name thy bname;
|
haftmann@25062
|
740 |
val (((sups, supconsts), (supsort, base_sort, mergeexpr)), elems_syn) =
|
haftmann@24748
|
741 |
prep_spec thy raw_supclasses raw_includes_elems;
|
wenzelm@24968
|
742 |
val other_consts = map (tap (Sign.the_const_type thy) o prep_param thy) raw_other_consts;
|
haftmann@24589
|
743 |
fun mk_inst class param_names cs =
|
haftmann@24589
|
744 |
Symtab.empty
|
haftmann@24589
|
745 |
|> fold2 (fn v => fn (c, ty) => Symtab.update (v, Const
|
haftmann@24589
|
746 |
(c, Term.map_type_tfree (fn (v, _) => TFree (v, [class])) ty))) param_names cs;
|
haftmann@25062
|
747 |
fun fork_syntax (Element.Fixes xs) =
|
haftmann@25062
|
748 |
fold_map (fn (c, ty, syn) => cons (c, syn) #> pair (c, ty, NoSyn)) xs
|
haftmann@25062
|
749 |
#>> Element.Fixes
|
haftmann@25062
|
750 |
| fork_syntax x = pair x;
|
haftmann@25062
|
751 |
val (elems, global_syn) = fold_map fork_syntax elems_syn [];
|
haftmann@25062
|
752 |
fun globalize (c, ty) =
|
haftmann@25062
|
753 |
((c, Term.map_type_tfree (K (TFree (Name.aT, base_sort))) ty),
|
haftmann@25062
|
754 |
(the_default NoSyn o AList.lookup (op =) global_syn) c);
|
haftmann@25038
|
755 |
fun extract_params thy =
|
haftmann@24218
|
756 |
let
|
haftmann@25062
|
757 |
val params = map fst (Locale.parameters_of thy class);
|
haftmann@24218
|
758 |
in
|
haftmann@25062
|
759 |
(params, (map globalize o snd o chop (length supconsts)) params)
|
haftmann@24218
|
760 |
end;
|
haftmann@25038
|
761 |
fun extract_assumes params thy cs =
|
haftmann@24218
|
762 |
let
|
haftmann@24218
|
763 |
val consts = supconsts @ (map (fst o fst) params ~~ cs);
|
haftmann@24218
|
764 |
fun subst (Free (c, ty)) =
|
haftmann@24218
|
765 |
Const ((fst o the o AList.lookup (op =) consts) c, ty)
|
haftmann@24218
|
766 |
| subst t = t;
|
haftmann@24218
|
767 |
fun prep_asm ((name, atts), ts) =
|
wenzelm@25024
|
768 |
((Sign.base_name name, map (Attrib.attribute_i thy) atts),
|
haftmann@24589
|
769 |
(map o map_aterms) subst ts);
|
haftmann@24218
|
770 |
in
|
haftmann@25038
|
771 |
Locale.global_asms_of thy class
|
haftmann@24218
|
772 |
|> map prep_asm
|
haftmann@24218
|
773 |
end;
|
haftmann@24218
|
774 |
in
|
haftmann@24218
|
775 |
thy
|
haftmann@24748
|
776 |
|> Locale.add_locale_i (SOME "") bname mergeexpr elems
|
haftmann@25038
|
777 |
|> snd
|
haftmann@25038
|
778 |
|> ProofContext.theory (`extract_params
|
haftmann@25062
|
779 |
#-> (fn (all_params, params) =>
|
wenzelm@24968
|
780 |
define_class_params (bname, supsort) params
|
haftmann@25038
|
781 |
(extract_assumes params) other_consts
|
haftmann@25038
|
782 |
#-> (fn (_, (consts, axioms)) =>
|
haftmann@25038
|
783 |
`(fn thy => class_intro thy class sups)
|
haftmann@24218
|
784 |
#-> (fn class_intro =>
|
haftmann@25062
|
785 |
PureThy.note_thmss_qualified "" (NameSpace.append class classN)
|
haftmann@25062
|
786 |
[((introN, []), [([class_intro], [])])]
|
haftmann@25062
|
787 |
#-> (fn [(_, [class_intro])] =>
|
haftmann@25038
|
788 |
add_class_data ((class, sups),
|
haftmann@25062
|
789 |
(map fst params ~~ consts, base_sort,
|
haftmann@25062
|
790 |
mk_inst class (map fst all_params) (map snd supconsts @ consts),
|
haftmann@25062
|
791 |
calculate_morphism class (supconsts @ (map (fst o fst) params ~~ consts)), class_intro))
|
haftmann@25038
|
792 |
#> class_interpretation class axioms []
|
haftmann@25062
|
793 |
)))))
|
haftmann@25038
|
794 |
|> pair class
|
haftmann@24218
|
795 |
end;
|
haftmann@24218
|
796 |
|
haftmann@24218
|
797 |
in
|
haftmann@24218
|
798 |
|
wenzelm@24968
|
799 |
val class_cmd = gen_class read_class_spec ((#1 o Term.dest_Const) oo Sign.read_const);
|
haftmann@24748
|
800 |
val class = gen_class check_class_spec (K I);
|
haftmann@24218
|
801 |
|
haftmann@24218
|
802 |
end; (*local*)
|
haftmann@24218
|
803 |
|
haftmann@24218
|
804 |
|
haftmann@24589
|
805 |
(* definition in class target *)
|
haftmann@24218
|
806 |
|
haftmann@25083
|
807 |
fun add_logical_const class ((c, mx), dict) thy =
|
haftmann@24218
|
808 |
let
|
wenzelm@25024
|
809 |
val prfx = class_prefix class;
|
wenzelm@25024
|
810 |
val thy' = thy |> Sign.add_path prfx;
|
haftmann@25062
|
811 |
val phi = morphism thy' class;
|
wenzelm@25024
|
812 |
|
haftmann@25062
|
813 |
val c' = Sign.full_name thy' c;
|
haftmann@25083
|
814 |
val dict' = (map_types Logic.unvarifyT o Morphism.term phi) dict;
|
haftmann@25083
|
815 |
val ty' = Term.fastype_of dict';
|
haftmann@25083
|
816 |
val ty'' = Type.strip_sorts ty';
|
haftmann@25083
|
817 |
val def_eq = Logic.mk_equals (Const (c', ty'), dict');
|
haftmann@25062
|
818 |
val c'' = NameSpace.full (Sign.naming_of thy' |> NameSpace.add_path prfx) c;
|
haftmann@24218
|
819 |
in
|
wenzelm@25024
|
820 |
thy'
|
haftmann@25038
|
821 |
|> Sign.hide_consts_i false [c'']
|
haftmann@25083
|
822 |
|> Sign.declare_const [] (c, ty'', mx) |> snd
|
haftmann@24218
|
823 |
|> Sign.parent_path
|
haftmann@24218
|
824 |
|> Sign.sticky_prefix prfx
|
haftmann@25083
|
825 |
|> Thm.add_def false (c, def_eq)
|
haftmann@25062
|
826 |
|>> Thm.symmetric
|
haftmann@25083
|
827 |
|-> (fn def => class_interpretation class [def] [Thm.prop_of def]
|
haftmann@25083
|
828 |
#> register_operation class ((c', dict), SOME (Thm.varifyT def)))
|
haftmann@24218
|
829 |
|> Sign.restore_naming thy
|
haftmann@25083
|
830 |
|> Sign.add_const_constraint (c', SOME ty')
|
haftmann@24914
|
831 |
|> pair c'
|
haftmann@24218
|
832 |
end;
|
haftmann@24218
|
833 |
|
haftmann@24901
|
834 |
|
haftmann@24901
|
835 |
(* abbreviation in class target *)
|
haftmann@24901
|
836 |
|
haftmann@25083
|
837 |
fun add_syntactic_const class prmode ((c, mx), rhs) thy =
|
haftmann@24836
|
838 |
let
|
wenzelm@25024
|
839 |
val prfx = class_prefix class;
|
haftmann@25062
|
840 |
val phi = morphism thy class;
|
haftmann@25062
|
841 |
|
haftmann@25038
|
842 |
val naming = Sign.naming_of thy |> NameSpace.add_path prfx |> NameSpace.add_path prfx;
|
haftmann@25038
|
843 |
(*FIXME*)
|
wenzelm@25024
|
844 |
val c' = NameSpace.full naming c;
|
haftmann@25062
|
845 |
val rhs' = (map_types Logic.unvarifyT o Morphism.term phi) rhs;
|
wenzelm@25020
|
846 |
val ty' = Term.fastype_of rhs';
|
haftmann@24836
|
847 |
in
|
haftmann@24836
|
848 |
thy
|
wenzelm@25024
|
849 |
|> Sign.notation true prmode [(Const (c', ty'), mx)]
|
haftmann@25062
|
850 |
|> register_operation class ((c', rhs), NONE)
|
haftmann@24914
|
851 |
|> pair c'
|
haftmann@24836
|
852 |
end;
|
haftmann@24836
|
853 |
|
haftmann@24218
|
854 |
end;
|