author | haftmann |
Fri, 24 Aug 2007 14:14:20 +0200 | |
changeset 24423 | ae9cd0e92423 |
parent 24304 | 69d40a562ba4 |
child 24435 | cabf8cd38258 |
permissions | -rw-r--r-- |
24218 | 1 |
(* Title: Pure/Isar/class.ML |
2 |
ID: $Id$ |
|
3 |
Author: Florian Haftmann, TU Muenchen |
|
4 |
||
5 |
Type classes derived from primitive axclasses and locales. |
|
6 |
*) |
|
7 |
||
8 |
signature CLASS = |
|
9 |
sig |
|
10 |
val fork_mixfix: bool -> string option -> mixfix -> mixfix * mixfix |
|
11 |
||
12 |
val axclass_cmd: bstring * xstring list |
|
13 |
-> ((bstring * Attrib.src list) * string list) list -> theory -> class * theory |
|
14 |
val class: bstring -> class list -> Element.context_i Locale.element list |
|
15 |
-> string list -> theory -> string * Proof.context |
|
16 |
val class_cmd: bstring -> string list -> Element.context Locale.element list |
|
17 |
-> string list -> theory -> string * Proof.context |
|
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
18 |
val class_of_locale: theory -> string -> class option |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
19 |
val add_const_in_class: string -> (string * term) * Syntax.mixfix |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
20 |
-> theory -> theory |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
21 |
|
24218 | 22 |
val instance_arity: arity list -> ((bstring * Attrib.src list) * term) list |
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
23 |
-> (thm list -> theory -> theory) |
24218 | 24 |
-> theory -> Proof.state |
25 |
val instance_arity_cmd: (bstring * string list * string) list |
|
26 |
-> ((bstring * Attrib.src list) * string) list |
|
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
27 |
-> (thm list -> theory -> theory) |
24218 | 28 |
-> theory -> Proof.state |
29 |
val prove_instance_arity: tactic -> arity list |
|
30 |
-> ((bstring * Attrib.src list) * term) list |
|
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
31 |
-> theory -> thm list * theory |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
32 |
val unoverload: theory -> conv |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
33 |
val overload: theory -> conv |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
34 |
val unoverload_const: theory -> string * typ -> string |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
35 |
val inst_const: theory -> string * string -> string |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
36 |
val param_const: theory -> string -> (string * string) option |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
37 |
val intro_classes_tac: thm list -> tactic |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
38 |
val default_intro_classes_tac: thm list -> tactic |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
39 |
|
24218 | 40 |
val instance_class: class * class -> theory -> Proof.state |
41 |
val instance_class_cmd: string * string -> theory -> Proof.state |
|
42 |
val instance_sort: class * sort -> theory -> Proof.state |
|
43 |
val instance_sort_cmd: string * string -> theory -> Proof.state |
|
44 |
val prove_instance_sort: tactic -> class * sort -> theory -> theory |
|
45 |
||
46 |
val print_classes: theory -> unit |
|
47 |
end; |
|
48 |
||
49 |
structure Class : CLASS = |
|
50 |
struct |
|
51 |
||
52 |
(** auxiliary **) |
|
53 |
||
54 |
fun fork_mixfix is_loc some_class mx = |
|
55 |
let |
|
56 |
val mx' = Syntax.unlocalize_mixfix mx; |
|
57 |
val mx_global = if is_some some_class orelse (is_loc andalso mx = mx') |
|
58 |
then NoSyn else mx'; |
|
59 |
val mx_local = if is_loc then mx else NoSyn; |
|
60 |
in (mx_global, mx_local) end; |
|
61 |
||
62 |
fun axclass_cmd (class, raw_superclasses) raw_specs thy = |
|
63 |
let |
|
64 |
val ctxt = ProofContext.init thy; |
|
65 |
val superclasses = map (Sign.read_class thy) raw_superclasses; |
|
66 |
val name_atts = map ((apsnd o map) (Attrib.attribute thy) o fst) raw_specs; |
|
67 |
val axiomss = ProofContext.read_propp (ctxt, map (map (rpair []) o snd) raw_specs) |
|
68 |
|> snd |
|
69 |
|> (map o map) fst; |
|
70 |
in AxClass.define_class (class, superclasses) [] (name_atts ~~ axiomss) thy end; |
|
71 |
||
72 |
||
73 |
(** axclasses with implicit parameter handling **) |
|
74 |
||
75 |
(* axclass instances *) |
|
76 |
||
77 |
local |
|
78 |
||
79 |
fun gen_instance mk_prop add_thm after_qed insts thy = |
|
80 |
let |
|
81 |
fun after_qed' results = |
|
82 |
ProofContext.theory ((fold o fold) add_thm results #> after_qed); |
|
83 |
in |
|
84 |
thy |
|
85 |
|> ProofContext.init |
|
86 |
|> Proof.theorem_i NONE after_qed' ((map (fn t => [(t, [])]) o maps (mk_prop thy)) insts) |
|
87 |
end; |
|
88 |
||
89 |
in |
|
90 |
||
91 |
val axclass_instance_arity = |
|
92 |
gen_instance (Logic.mk_arities oo Sign.cert_arity) AxClass.add_arity; |
|
93 |
val axclass_instance_sort = |
|
94 |
gen_instance (single oo (Logic.mk_classrel oo AxClass.cert_classrel)) |
|
95 |
AxClass.add_classrel I o single; |
|
96 |
||
97 |
end; (*local*) |
|
98 |
||
99 |
||
100 |
(* introducing axclasses with implicit parameter handling *) |
|
101 |
||
102 |
fun axclass_params (name, raw_superclasses) raw_consts raw_dep_axioms other_consts thy = |
|
103 |
let |
|
104 |
val superclasses = map (Sign.certify_class thy) raw_superclasses; |
|
105 |
val consts = (map o apfst o apsnd) (Sign.certify_typ thy) raw_consts; |
|
106 |
val prefix = Logic.const_of_class name; |
|
107 |
fun mk_const_name c = NameSpace.map_base (NameSpace.append prefix) |
|
108 |
(Sign.full_name thy c); |
|
109 |
fun add_const ((c, ty), syn) = |
|
110 |
Sign.add_consts_authentic [(c, ty, syn)] |
|
111 |
#> pair (mk_const_name c, ty); |
|
112 |
fun mk_axioms cs thy = |
|
113 |
raw_dep_axioms thy cs |
|
114 |
|> (map o apsnd o map) (Sign.cert_prop thy) |
|
115 |
|> rpair thy; |
|
116 |
fun add_constraint class (c, ty) = |
|
117 |
Sign.add_const_constraint_i (c, SOME |
|
118 |
(Term.map_type_tfree (fn (v, _) => TFree (v, [class])) ty)); |
|
119 |
in |
|
120 |
thy |
|
121 |
|> Theory.add_path prefix |
|
122 |
|> fold_map add_const consts |
|
123 |
||> Theory.restore_naming thy |
|
124 |
|-> (fn cs => mk_axioms cs |
|
125 |
#-> (fn axioms_prop => AxClass.define_class (name, superclasses) |
|
126 |
(map fst cs @ other_consts) axioms_prop |
|
127 |
#-> (fn class => `(fn thy => AxClass.get_definition thy class) |
|
128 |
#-> (fn {intro, axioms, ...} => fold (add_constraint class) cs |
|
129 |
#> pair (class, ((intro, (map Thm.prop_of axioms, axioms)), cs)))))) |
|
130 |
end; |
|
131 |
||
132 |
||
24304 | 133 |
(* explicit constants for overloaded definitions *) |
134 |
||
135 |
structure InstData = TheoryDataFun |
|
136 |
( |
|
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
137 |
type T = (string * thm) Symtab.table Symtab.table * (string * string) Symtab.table; |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
138 |
(*constant name ~> type constructor ~> (constant name, equation), |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
139 |
constant name ~> (constant name, type constructor)*) |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
140 |
val empty = (Symtab.empty, Symtab.empty); |
24304 | 141 |
val copy = I; |
142 |
val extend = I; |
|
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
143 |
fun merge _ ((taba1, tabb1), (taba2, tabb2)) = |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
144 |
(Symtab.join (K (Symtab.merge (K true))) (taba1, taba2), |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
145 |
Symtab.merge (K true) (tabb1, tabb2)); |
24304 | 146 |
); |
147 |
||
148 |
fun inst_thms f thy = |
|
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
149 |
(Symtab.fold (Symtab.fold (cons o f o snd o snd) o snd) o fst) (InstData.get thy) []; |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
150 |
fun add_inst (c, tyco) inst = (InstData.map o apfst o Symtab.map_default (c, Symtab.empty)) |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
151 |
(Symtab.update_new (tyco, inst)) |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
152 |
#> (InstData.map o apsnd) (Symtab.update_new (fst inst, (c, tyco))); |
24304 | 153 |
|
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
154 |
fun unoverload thy thm = MetaSimplifier.rewrite false (inst_thms I thy) thm; |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
155 |
fun overload thy thm = MetaSimplifier.rewrite false (inst_thms symmetric thy) thm; |
24304 | 156 |
|
157 |
fun inst_const thy (c, tyco) = |
|
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
158 |
(fst o the o Symtab.lookup ((the o Symtab.lookup (fst (InstData.get thy))) c)) tyco; |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
159 |
fun unoverload_const thy (c_ty as (c, _)) = |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
160 |
case AxClass.class_of_param thy c |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
161 |
of SOME class => (case Sign.const_typargs thy c_ty |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
162 |
of [Type (tyco, _)] => |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
163 |
(case Symtab.lookup ((the o Symtab.lookup (fst (InstData.get thy))) c) tyco |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
164 |
of SOME (c, _) => c |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
165 |
| NONE => c) |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
166 |
| [_] => c) |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
167 |
| NONE => c; |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
168 |
|
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
169 |
val param_const = Symtab.lookup o snd o InstData.get; |
24304 | 170 |
|
171 |
fun add_inst_def (class, tyco) (c, ty) thy = |
|
172 |
let |
|
173 |
val tyco_base = NameSpace.base tyco; |
|
174 |
val name_inst = NameSpace.base class ^ "_" ^ tyco_base ^ "_inst"; |
|
175 |
val c_inst_base = NameSpace.base c ^ "_" ^ tyco_base; |
|
176 |
in |
|
177 |
thy |
|
178 |
|> Sign.sticky_prefix name_inst |
|
179 |
|> Sign.add_consts_i [(c_inst_base, ty, Syntax.NoSyn)] |
|
180 |
|> `(fn thy => Sign.full_name thy c_inst_base) |
|
181 |
|-> (fn c_inst => PureThy.add_defs_i true |
|
182 |
[((Thm.def_name c_inst_base, Logic.mk_equals (Const (c_inst, ty), Const (c, ty))), [])] |
|
183 |
#-> (fn [def] => add_inst (c, tyco) (c_inst, symmetric def)) |
|
184 |
#> Sign.restore_naming thy) |
|
185 |
end; |
|
186 |
||
187 |
fun add_inst_def' (class, tyco) (c, ty) thy = |
|
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
188 |
if case Symtab.lookup (fst (InstData.get thy)) c |
24304 | 189 |
of NONE => true |
190 |
| SOME tab => is_none (Symtab.lookup tab tyco) |
|
191 |
then add_inst_def (class, tyco) (c, Logic.unvarifyT ty) thy |
|
192 |
else thy; |
|
193 |
||
194 |
fun add_def ((class, tyco), ((name, prop), atts)) thy = |
|
195 |
let |
|
196 |
val ((lhs as Const (c, ty), args), rhs) = (apfst Term.strip_comb o Logic.dest_equals) prop; |
|
197 |
fun add_inst' def ([], (Const (c_inst, ty))) = |
|
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
198 |
if forall (fn TFree _ => true | _ => false) (Sign.const_typargs thy (c_inst, ty)) |
24304 | 199 |
then add_inst (c, tyco) (c_inst, def) |
200 |
else add_inst_def (class, tyco) (c, ty) |
|
201 |
| add_inst' _ t = add_inst_def (class, tyco) (c, ty); |
|
202 |
in |
|
203 |
thy |
|
204 |
|> PureThy.add_defs_i true [((name, prop), map (Attrib.attribute thy) atts)] |
|
205 |
|-> (fn [def] => add_inst' def (args, rhs) #> pair def) |
|
206 |
end; |
|
207 |
||
208 |
||
24218 | 209 |
(* instances with implicit parameter handling *) |
210 |
||
211 |
local |
|
212 |
||
213 |
fun gen_read_def thy prep_att read_def ((raw_name, raw_atts), raw_t) = |
|
214 |
let |
|
215 |
val (_, t) = read_def thy (raw_name, raw_t); |
|
216 |
val ((c, ty), _) = Sign.cert_def (Sign.pp thy) t; |
|
217 |
val atts = map (prep_att thy) raw_atts; |
|
218 |
val insts = Consts.typargs (Sign.consts_of thy) (c, ty); |
|
219 |
val name = case raw_name |
|
220 |
of "" => NONE |
|
221 |
| _ => SOME raw_name; |
|
222 |
in (c, (insts, ((name, t), atts))) end; |
|
223 |
||
224 |
fun read_def_cmd thy = gen_read_def thy Attrib.intern_src Theory.read_axm; |
|
225 |
fun read_def thy = gen_read_def thy (K I) (K I); |
|
226 |
||
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
227 |
fun gen_instance_arity prep_arity read_def do_proof raw_arities raw_defs after_qed theory = |
24218 | 228 |
let |
229 |
val arities = map (prep_arity theory) raw_arities; |
|
230 |
val _ = if null arities then error "at least one arity must be given" else (); |
|
231 |
val _ = case (duplicates (op =) o map #1) arities |
|
232 |
of [] => () |
|
233 |
| dupl_tycos => error ("type constructors occur more than once in arities: " |
|
234 |
^ (commas o map quote) dupl_tycos); |
|
235 |
val super_sort = (Graph.all_succs o #classes o Sorts.rep_algebra o Sign.classes_of) theory |
|
236 |
fun get_consts_class tyco ty class = |
|
237 |
let |
|
238 |
val cs = (these o Option.map snd o try (AxClass.params_of_class theory)) class; |
|
239 |
val subst_ty = map_type_tfree (K ty); |
|
240 |
in |
|
24304 | 241 |
map (fn (c, ty) => (c, ((class, tyco), subst_ty ty))) cs |
24218 | 242 |
end; |
243 |
fun get_consts_sort (tyco, asorts, sort) = |
|
244 |
let |
|
245 |
val ty = Type (tyco, map (fn (v, sort) => TVar ((v, 0), sort)) (Name.names Name.context "'a" asorts)) |
|
246 |
in maps (get_consts_class tyco ty) (super_sort sort) end; |
|
247 |
val cs = maps get_consts_sort arities; |
|
248 |
fun mk_typnorm thy (ty, ty_sc) = |
|
249 |
case try (Sign.typ_match thy (Logic.varifyT ty_sc, ty)) Vartab.empty |
|
250 |
of SOME env => SOME (Logic.varifyT #> Envir.typ_subst_TVars env #> Logic.unvarifyT) |
|
251 |
| NONE => NONE; |
|
252 |
fun read_defs defs cs thy_read = |
|
253 |
let |
|
254 |
fun read raw_def cs = |
|
255 |
let |
|
256 |
val (c, (inst, ((name_opt, t), atts))) = read_def thy_read raw_def; |
|
257 |
val ty = Consts.instance (Sign.consts_of thy_read) (c, inst); |
|
24304 | 258 |
val ((class, tyco), ty') = case AList.lookup (op =) cs c |
24218 | 259 |
of NONE => error ("illegal definition for constant " ^ quote c) |
260 |
| SOME class_ty => class_ty; |
|
261 |
val name = case name_opt |
|
262 |
of NONE => Thm.def_name (Logic.name_arity (tyco, [], c)) |
|
263 |
| SOME name => name; |
|
264 |
val t' = case mk_typnorm thy_read (ty', ty) |
|
265 |
of NONE => error ("illegal definition for constant " ^ |
|
266 |
quote (c ^ "::" ^ setmp show_sorts true |
|
267 |
(Sign.string_of_typ thy_read) ty)) |
|
268 |
| SOME norm => map_types norm t |
|
269 |
in (((class, tyco), ((name, t'), atts)), AList.delete (op =) c cs) end; |
|
270 |
in fold_map read defs cs end; |
|
24304 | 271 |
val (defs, other_cs) = read_defs raw_defs cs |
24218 | 272 |
(fold Sign.primitive_arity arities (Theory.copy theory)); |
273 |
fun get_remove_contraint c thy = |
|
274 |
let |
|
275 |
val ty = Sign.the_const_constraint thy c; |
|
276 |
in |
|
277 |
thy |
|
278 |
|> Sign.add_const_constraint_i (c, NONE) |
|
279 |
|> pair (c, Logic.unvarifyT ty) |
|
280 |
end; |
|
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
281 |
fun after_qed' cs defs = |
24304 | 282 |
fold Sign.add_const_constraint_i (map (apsnd SOME) cs) |
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
283 |
#> after_qed defs; |
24218 | 284 |
in |
285 |
theory |
|
286 |
|> fold_map get_remove_contraint (map fst cs |> distinct (op =)) |
|
24304 | 287 |
||>> fold_map add_def defs |
288 |
||> fold (fn (c, ((class, tyco), ty)) => add_inst_def' (class, tyco) (c, ty)) other_cs |
|
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
289 |
|-> (fn (cs, defs) => do_proof (after_qed' cs defs) arities defs) |
24218 | 290 |
end; |
291 |
||
292 |
fun instance_arity_cmd' do_proof = gen_instance_arity Sign.read_arity read_def_cmd do_proof; |
|
293 |
fun instance_arity' do_proof = gen_instance_arity Sign.cert_arity read_def do_proof; |
|
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
294 |
fun tactic_proof tac f arities defs = |
24218 | 295 |
fold (fn arity => AxClass.prove_arity arity tac) arities |
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
296 |
#> f |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
297 |
#> pair defs; |
24218 | 298 |
|
299 |
in |
|
300 |
||
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
301 |
val instance_arity_cmd = instance_arity_cmd' |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
302 |
(fn f => fn arities => fn defs => axclass_instance_arity f arities); |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
303 |
val instance_arity = instance_arity' |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
304 |
(fn f => fn arities => fn defs => axclass_instance_arity f arities); |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
305 |
fun prove_instance_arity tac arities defs = |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24304
diff
changeset
|
306 |
instance_arity' (tactic_proof tac) arities defs (K I); |
24218 | 307 |
|
308 |
end; (*local*) |
|
309 |
||
310 |
||
311 |
||
312 |
(** combining locales and axclasses **) |
|
313 |
||
314 |
(* theory data *) |
|
315 |
||
316 |
datatype class_data = ClassData of { |
|
317 |
locale: string, |
|
318 |
consts: (string * string) list |
|
319 |
(*locale parameter ~> toplevel theory constant*), |
|
320 |
v: string option, |
|
321 |
intro: thm |
|
322 |
} * thm list (*derived defs*); |
|
323 |
||
324 |
fun rep_classdata (ClassData c) = c; |
|
325 |
||
326 |
fun merge_pair f1 f2 ((x1, y1), (x2, y2)) = (f1 (x1, x2), f2 (y1, y2)); |
|
327 |
||
328 |
structure ClassData = TheoryDataFun |
|
329 |
( |
|
330 |
type T = class_data Graph.T * class Symtab.table (*locale name ~> class name*); |
|
331 |
val empty = (Graph.empty, Symtab.empty); |
|
332 |
val copy = I; |
|
333 |
val extend = I; |
|
334 |
fun merge _ = merge_pair (Graph.merge (K true)) (Symtab.merge (K true)); |
|
335 |
); |
|
336 |
||
337 |
||
338 |
(* queries *) |
|
339 |
||
340 |
val lookup_class_data = Option.map rep_classdata oo try o Graph.get_node o fst o ClassData.get; |
|
341 |
fun class_of_locale thy = Symtab.lookup ((snd o ClassData.get) thy); |
|
342 |
||
343 |
fun the_class_data thy class = |
|
344 |
case lookup_class_data thy class |
|
345 |
of NONE => error ("undeclared class " ^ quote class) |
|
346 |
| SOME data => data; |
|
347 |
||
348 |
val ancestry = Graph.all_succs o fst o ClassData.get; |
|
349 |
||
350 |
fun param_map thy = |
|
351 |
let |
|
352 |
fun params class = |
|
353 |
let |
|
354 |
val const_typs = (#params o AxClass.get_definition thy) class; |
|
355 |
val const_names = (#consts o fst o the_class_data thy) class; |
|
356 |
in |
|
357 |
(map o apsnd) (fn c => (c, (the o AList.lookup (op =) const_typs) c)) const_names |
|
358 |
end; |
|
359 |
in maps params o ancestry thy end; |
|
360 |
||
361 |
fun these_defs thy = maps (these o Option.map snd o lookup_class_data thy) o ancestry thy; |
|
362 |
||
363 |
fun these_intros thy = |
|
364 |
Graph.fold (fn (_, (data, _)) => insert Thm.eq_thm ((#intro o fst o rep_classdata) data)) |
|
365 |
((fst o ClassData.get) thy) []; |
|
366 |
||
367 |
fun print_classes thy = |
|
368 |
let |
|
369 |
val algebra = Sign.classes_of thy; |
|
370 |
val arities = |
|
371 |
Symtab.empty |
|
372 |
|> Symtab.fold (fn (tyco, arities) => fold (fn (class, _) => |
|
373 |
Symtab.map_default (class, []) (insert (op =) tyco)) arities) |
|
374 |
((#arities o Sorts.rep_algebra) algebra); |
|
375 |
val the_arities = these o Symtab.lookup arities; |
|
376 |
fun mk_arity class tyco = |
|
377 |
let |
|
378 |
val Ss = Sorts.mg_domain algebra tyco [class]; |
|
379 |
in Sign.pretty_arity thy (tyco, Ss, [class]) end; |
|
380 |
fun mk_param (c, ty) = Pretty.str (Sign.extern_const thy c ^ " :: " |
|
381 |
^ setmp show_sorts false (Sign.string_of_typ thy o Type.strip_sorts) ty); |
|
382 |
fun mk_entry class = (Pretty.block o Pretty.fbreaks o map_filter I) [ |
|
383 |
(SOME o Pretty.str) ("class " ^ class ^ ":"), |
|
384 |
(SOME o Pretty.block) [Pretty.str "supersort: ", |
|
385 |
(Sign.pretty_sort thy o Sign.certify_sort thy o Sign.super_classes thy) class], |
|
386 |
Option.map (Pretty.str o prefix "locale: " o #locale o fst) (lookup_class_data thy class), |
|
387 |
((fn [] => NONE | ps => (SOME o Pretty.block o Pretty.fbreaks) (Pretty.str "parameters:" :: ps)) o map mk_param |
|
388 |
o these o Option.map #params o try (AxClass.get_definition thy)) class, |
|
389 |
(SOME o Pretty.block o Pretty.breaks) [ |
|
390 |
Pretty.str "instances:", |
|
391 |
Pretty.list "" "" (map (mk_arity class) (the_arities class)) |
|
392 |
] |
|
393 |
] |
|
394 |
in |
|
395 |
(Pretty.writeln o Pretty.chunks o separate (Pretty.str "") o map mk_entry o Sorts.all_classes) |
|
396 |
algebra |
|
397 |
end; |
|
398 |
||
399 |
||
400 |
(* updaters *) |
|
401 |
||
402 |
fun add_class_data ((class, superclasses), (locale, consts, v, intro)) = |
|
403 |
ClassData.map (fn (gr, tab) => ( |
|
404 |
gr |
|
405 |
|> Graph.new_node (class, ClassData ({ locale = locale, consts = consts, |
|
406 |
v = v, intro = intro }, [])) |
|
407 |
|> fold (curry Graph.add_edge class) superclasses, |
|
408 |
tab |
|
409 |
|> Symtab.update (locale, class) |
|
410 |
)); |
|
411 |
||
412 |
fun add_class_const_thm (class, thm) = (ClassData.map o apfst o Graph.map_node class) |
|
413 |
(fn ClassData (data, thms) => ClassData (data, thm :: thms)); |
|
414 |
||
415 |
(* tactics and methods *) |
|
416 |
||
417 |
fun intro_classes_tac facts st = |
|
418 |
let |
|
419 |
val thy = Thm.theory_of_thm st; |
|
420 |
val classes = Sign.all_classes thy; |
|
421 |
val class_trivs = map (Thm.class_triv thy) classes; |
|
422 |
val class_intros = these_intros thy; |
|
423 |
fun add_axclass_intro class = |
|
424 |
case try (AxClass.get_definition thy) class of SOME {intro, ...} => cons intro | _ => I; |
|
425 |
val axclass_intros = fold add_axclass_intro classes []; |
|
426 |
in |
|
427 |
st |
|
428 |
|> ((ALLGOALS (Method.insert_tac facts THEN' |
|
429 |
REPEAT_ALL_NEW (resolve_tac (class_trivs @ class_intros @ axclass_intros)))) |
|
430 |
THEN Tactic.distinct_subgoals_tac) |
|
431 |
end; |
|
432 |
||
433 |
fun default_intro_classes_tac [] = intro_classes_tac [] |
|
434 |
| default_intro_classes_tac _ = Tactical.no_tac; (*no error message!*) |
|
435 |
||
436 |
fun default_tac rules ctxt facts = |
|
437 |
HEADGOAL (Method.some_rule_tac rules ctxt facts) ORELSE |
|
438 |
default_intro_classes_tac facts; |
|
439 |
||
440 |
val _ = Context.add_setup (Method.add_methods |
|
441 |
[("intro_classes", Method.no_args (Method.METHOD intro_classes_tac), |
|
442 |
"back-chain introduction rules of classes"), |
|
443 |
("default", Method.thms_ctxt_args (Method.METHOD oo default_tac), |
|
444 |
"apply some intro/elim rule")]); |
|
445 |
||
446 |
||
447 |
(* tactical interfaces to locale commands *) |
|
448 |
||
449 |
fun prove_interpretation tac prfx_atts expr insts thy = |
|
450 |
thy |
|
451 |
|> Locale.interpretation_i I prfx_atts expr insts |
|
452 |
|> Proof.global_terminal_proof (Method.Basic (K (Method.SIMPLE_METHOD tac), Position.none), NONE) |
|
453 |
|> ProofContext.theory_of; |
|
454 |
||
455 |
fun prove_interpretation_in tac after_qed (name, expr) thy = |
|
456 |
thy |
|
457 |
|> Locale.interpretation_in_locale (ProofContext.theory after_qed) (name, expr) |
|
458 |
|> Proof.global_terminal_proof (Method.Basic (K (Method.SIMPLE_METHOD tac), Position.none), NONE) |
|
459 |
|> ProofContext.theory_of; |
|
460 |
||
461 |
||
462 |
(* constructing class introduction and other rules from axclass and locale rules *) |
|
463 |
||
464 |
fun mk_instT class = Symtab.empty |
|
465 |
|> Symtab.update (AxClass.param_tyvarname, TFree (AxClass.param_tyvarname, [class])); |
|
466 |
||
467 |
fun mk_inst class param_names cs = |
|
468 |
Symtab.empty |
|
469 |
|> fold2 (fn v => fn (c, ty) => Symtab.update (v, Const |
|
470 |
(c, Term.map_type_tfree (fn (v, _) => TFree (v, [class])) ty))) param_names cs; |
|
471 |
||
472 |
fun OF_LAST thm1 thm2 = |
|
473 |
let |
|
474 |
val n = (length o Logic.strip_imp_prems o prop_of) thm2; |
|
475 |
in (thm1 RSN (n, thm2)) end; |
|
476 |
||
477 |
fun strip_all_ofclass thy sort = |
|
478 |
let |
|
479 |
val typ = TVar ((AxClass.param_tyvarname, 0), sort); |
|
480 |
fun prem_inclass t = |
|
481 |
case Logic.strip_imp_prems t |
|
482 |
of ofcls :: _ => try Logic.dest_inclass ofcls |
|
483 |
| [] => NONE; |
|
484 |
fun strip_ofclass class thm = |
|
485 |
thm OF (fst o AxClass.of_sort thy (typ, [class])) AxClass.cache; |
|
486 |
fun strip thm = case (prem_inclass o Thm.prop_of) thm |
|
487 |
of SOME (_, class) => thm |> strip_ofclass class |> strip |
|
488 |
| NONE => thm; |
|
489 |
in strip end; |
|
490 |
||
491 |
fun class_intro thy locale class sups = |
|
492 |
let |
|
493 |
fun class_elim class = |
|
494 |
case (map Drule.unconstrainTs o #axioms o AxClass.get_definition thy) class |
|
495 |
of [thm] => SOME thm |
|
496 |
| [] => NONE; |
|
497 |
val pred_intro = case Locale.intros thy locale |
|
498 |
of ([ax_intro], [intro]) => intro |> OF_LAST ax_intro |> SOME |
|
499 |
| ([intro], []) => SOME intro |
|
500 |
| ([], [intro]) => SOME intro |
|
501 |
| _ => NONE; |
|
502 |
val pred_intro' = pred_intro |
|
503 |
|> Option.map (fn intro => intro OF map_filter class_elim sups); |
|
504 |
val class_intro = (#intro o AxClass.get_definition thy) class; |
|
505 |
val raw_intro = case pred_intro' |
|
506 |
of SOME pred_intro => class_intro |> OF_LAST pred_intro |
|
507 |
| NONE => class_intro; |
|
508 |
val sort = Sign.super_classes thy class; |
|
509 |
val typ = TVar ((AxClass.param_tyvarname, 0), sort); |
|
510 |
val defs = these_defs thy sups; |
|
511 |
in |
|
512 |
raw_intro |
|
513 |
|> Drule.instantiate' [SOME (Thm.ctyp_of thy typ)] [] |
|
514 |
|> strip_all_ofclass thy sort |
|
515 |
|> Thm.strip_shyps |
|
516 |
|> MetaSimplifier.rewrite_rule defs |
|
517 |
|> Drule.unconstrainTs |
|
518 |
end; |
|
519 |
||
520 |
fun interpretation_in_rule thy (class1, class2) = |
|
521 |
let |
|
522 |
val (params, consts) = split_list (param_map thy [class1]); |
|
523 |
(*FIXME also remember this at add_class*) |
|
524 |
fun mk_axioms class = |
|
525 |
let |
|
526 |
val name_locale = (#locale o fst o the_class_data thy) class; |
|
527 |
val inst = mk_inst class params consts; |
|
528 |
in |
|
529 |
Locale.global_asms_of thy name_locale |
|
530 |
|> maps snd |
|
531 |
|> (map o map_aterms) (fn Free (s, _) => (the o Symtab.lookup inst) s | t => t) |
|
532 |
|> (map o map_types o map_atyps) (fn TFree _ => TFree ("'a", [class1]) | T => T) |
|
533 |
|> map (ObjectLogic.ensure_propT thy) |
|
534 |
end; |
|
535 |
val (prems, concls) = pairself mk_axioms (class1, class2); |
|
536 |
in |
|
537 |
Goal.prove_global thy [] prems (Logic.mk_conjunction_list concls) |
|
538 |
(Locale.intro_locales_tac true (ProofContext.init thy)) |
|
539 |
end; |
|
540 |
||
541 |
||
542 |
(* classes *) |
|
543 |
||
544 |
local |
|
545 |
||
546 |
fun read_param thy raw_t = |
|
547 |
let |
|
548 |
val t = Sign.read_term thy raw_t |
|
549 |
in case try dest_Const t |
|
550 |
of SOME (c, _) => c |
|
551 |
| NONE => error ("Not a constant: " ^ Sign.string_of_term thy t) |
|
552 |
end; |
|
553 |
||
554 |
fun gen_class add_locale prep_class prep_param bname |
|
555 |
raw_supclasses raw_elems raw_other_consts thy = |
|
556 |
let |
|
557 |
(*FIXME need proper concept for reading locale statements*) |
|
558 |
fun subst_classtyvar (_, _) = |
|
559 |
TFree (AxClass.param_tyvarname, []) |
|
560 |
| subst_classtyvar (v, sort) = |
|
561 |
error ("Sort constraint illegal in type class, for type variable " ^ v ^ "::" ^ Sign.string_of_sort thy sort); |
|
562 |
(*val subst_classtyvars = Element.map_ctxt {name = I, var = I, term = I, |
|
563 |
typ = Term.map_type_tfree subst_classtyvar, fact = I, attrib = I};*) |
|
564 |
val other_consts = map (prep_param thy) raw_other_consts; |
|
565 |
val (elems, includes) = fold_rev (fn Locale.Elem e => apfst (cons e) |
|
566 |
| Locale.Expr i => apsnd (cons i)) raw_elems ([], []); |
|
567 |
val supclasses = map (prep_class thy) raw_supclasses; |
|
568 |
val sups = filter (is_some o lookup_class_data thy) supclasses |
|
569 |
|> Sign.certify_sort thy; |
|
570 |
val supsort = Sign.certify_sort thy supclasses; |
|
571 |
val suplocales = map (Locale.Locale o #locale o fst o the_class_data thy) sups; |
|
572 |
val supexpr = Locale.Merge (suplocales @ includes); |
|
573 |
val supparams = (map fst o Locale.parameters_of_expr thy) |
|
574 |
(Locale.Merge suplocales); |
|
575 |
val supconsts = AList.make (the o AList.lookup (op =) (param_map thy sups)) |
|
576 |
(map fst supparams); |
|
577 |
(*val elems_constrains = map |
|
578 |
(Element.Constrains o apsnd (Term.map_type_tfree subst_classtyvar)) supparams;*) |
|
579 |
fun mk_tyvar (_, sort) = TFree (AxClass.param_tyvarname, |
|
580 |
if Sign.subsort thy (supsort, sort) then sort else error |
|
581 |
("Sort " ^ Sign.string_of_sort thy sort |
|
582 |
^ " is less general than permitted least general sort " |
|
583 |
^ Sign.string_of_sort thy supsort)); |
|
584 |
fun extract_params thy name_locale = |
|
585 |
let |
|
586 |
val params = Locale.parameters_of thy name_locale; |
|
587 |
val v = case (maps typ_tfrees o map (snd o fst)) params |
|
588 |
of (v, _) :: _ => SOME v |
|
589 |
| _ => NONE; |
|
590 |
in |
|
591 |
(v, (map (fst o fst) params, params |
|
592 |
|> (map o apfst o apsnd o Term.map_type_tfree) mk_tyvar |
|
593 |
|> (map o apsnd) (fork_mixfix true NONE #> fst) |
|
594 |
|> chop (length supconsts) |
|
595 |
|> snd)) |
|
596 |
end; |
|
597 |
fun extract_assumes name_locale params thy cs = |
|
598 |
let |
|
599 |
val consts = supconsts @ (map (fst o fst) params ~~ cs); |
|
600 |
fun subst (Free (c, ty)) = |
|
601 |
Const ((fst o the o AList.lookup (op =) consts) c, ty) |
|
602 |
| subst t = t; |
|
603 |
val super_defs = these_defs thy sups; |
|
604 |
fun prep_asm ((name, atts), ts) = |
|
605 |
((NameSpace.base name, map (Attrib.attribute thy) atts), |
|
606 |
(map o map_aterms) ((*MetaSimplifier.rewrite_term thy super_defs [] o *)subst) ts); |
|
607 |
in |
|
608 |
Locale.global_asms_of thy name_locale |
|
609 |
|> map prep_asm |
|
610 |
end; |
|
611 |
fun note_intro name_axclass class_intro = |
|
612 |
PureThy.note_thmss_qualified "" ((Logic.const_of_class o NameSpace.base) name_axclass) |
|
613 |
[(("intro", []), [([class_intro], [])])] |
|
614 |
#> snd |
|
615 |
in |
|
616 |
thy |
|
617 |
|> add_locale (SOME "") bname supexpr ((*elems_constrains @*) elems) |
|
618 |
|-> (fn name_locale => ProofContext.theory_result ( |
|
619 |
`(fn thy => extract_params thy name_locale) |
|
620 |
#-> (fn (v, (param_names, params)) => |
|
621 |
axclass_params (bname, supsort) params (extract_assumes name_locale params) other_consts |
|
622 |
#-> (fn (name_axclass, ((_, (ax_terms, ax_axioms)), consts)) => |
|
623 |
`(fn thy => class_intro thy name_locale name_axclass sups) |
|
624 |
#-> (fn class_intro => |
|
625 |
add_class_data ((name_axclass, sups), |
|
626 |
(name_locale, map (fst o fst) params ~~ map fst consts, v, |
|
627 |
class_intro)) |
|
628 |
(*FIXME: class_data should already contain data relevant |
|
629 |
for interpretation; use this later for class target*) |
|
630 |
(*FIXME: general export_fixes which may be parametrized |
|
631 |
with pieces of an emerging class*) |
|
632 |
#> note_intro name_axclass class_intro |
|
633 |
#> prove_interpretation ((ALLGOALS o ProofContext.fact_tac) ax_axioms) |
|
634 |
((false, Logic.const_of_class bname), []) (Locale.Locale name_locale) |
|
635 |
((mk_instT name_axclass, mk_inst name_axclass param_names (map snd supconsts @ consts)), []) |
|
636 |
#> pair name_axclass |
|
637 |
))))) |
|
638 |
end; |
|
639 |
||
640 |
in |
|
641 |
||
642 |
val class_cmd = gen_class Locale.add_locale Sign.intern_class read_param; |
|
643 |
val class = gen_class Locale.add_locale_i Sign.certify_class (K I); |
|
644 |
||
645 |
end; (*local*) |
|
646 |
||
647 |
local |
|
648 |
||
649 |
fun instance_subclass (class1, class2) thy = |
|
650 |
let |
|
651 |
val interp = interpretation_in_rule thy (class1, class2); |
|
652 |
val ax = #axioms (AxClass.get_definition thy class1); |
|
653 |
val intro = #intro (AxClass.get_definition thy class2) |
|
654 |
|> Drule.instantiate' [SOME (Thm.ctyp_of thy |
|
655 |
(TVar ((AxClass.param_tyvarname, 0), [class1])))] []; |
|
656 |
val thm = |
|
657 |
intro |
|
658 |
|> OF_LAST (interp OF ax) |
|
659 |
|> strip_all_ofclass thy (Sign.super_classes thy class2); |
|
660 |
in |
|
661 |
thy |> AxClass.add_classrel thm |
|
662 |
end; |
|
663 |
||
664 |
fun instance_subsort (class, sort) thy = |
|
665 |
let |
|
666 |
val super_sort = (Graph.all_succs o #classes o Sorts.rep_algebra |
|
667 |
o Sign.classes_of) thy sort; |
|
668 |
val classes = filter_out (fn class' => Sign.subsort thy ([class], [class'])) |
|
669 |
(rev super_sort); |
|
670 |
in |
|
671 |
thy |> fold (curry instance_subclass class) classes |
|
672 |
end; |
|
673 |
||
674 |
fun instance_sort' do_proof (class, sort) theory = |
|
675 |
let |
|
676 |
val loc_name = (#locale o fst o the_class_data theory) class; |
|
677 |
val loc_expr = |
|
678 |
(Locale.Merge o map (Locale.Locale o #locale o fst o the_class_data theory)) sort; |
|
679 |
in |
|
680 |
theory |
|
681 |
|> do_proof (instance_subsort (class, sort)) (loc_name, loc_expr) |
|
682 |
end; |
|
683 |
||
684 |
fun gen_instance_sort prep_class prep_sort (raw_class, raw_sort) theory = |
|
685 |
let |
|
686 |
val class = prep_class theory raw_class; |
|
687 |
val sort = prep_sort theory raw_sort; |
|
688 |
in |
|
689 |
theory |
|
690 |
|> instance_sort' (Locale.interpretation_in_locale o ProofContext.theory) (class, sort) |
|
691 |
end; |
|
692 |
||
693 |
fun gen_instance_class prep_class (raw_class, raw_superclass) theory = |
|
694 |
let |
|
695 |
val class = prep_class theory raw_class; |
|
696 |
val superclass = prep_class theory raw_superclass; |
|
697 |
in |
|
698 |
theory |
|
699 |
|> axclass_instance_sort (class, superclass) |
|
700 |
end; |
|
701 |
||
702 |
in |
|
703 |
||
24276 | 704 |
val instance_sort_cmd = gen_instance_sort Sign.read_class Syntax.global_read_sort; |
24218 | 705 |
val instance_sort = gen_instance_sort Sign.certify_class Sign.certify_sort; |
706 |
val prove_instance_sort = instance_sort' o prove_interpretation_in; |
|
707 |
val instance_class_cmd = gen_instance_class Sign.read_class; |
|
708 |
val instance_class = gen_instance_class Sign.certify_class; |
|
709 |
||
710 |
end; (*local*) |
|
711 |
||
712 |
||
713 |
(** class target **) |
|
714 |
||
715 |
fun export_fixes thy class = |
|
716 |
let |
|
717 |
val v = (#v o fst o the_class_data thy) class; |
|
718 |
val constrain_sort = curry (Sorts.inter_sort (Sign.classes_of thy)) [class]; |
|
719 |
val subst_typ = Term.map_type_tfree (fn var as (w, sort) => |
|
720 |
if SOME w = v then TFree (w, constrain_sort sort) else TFree var); |
|
721 |
val consts = param_map thy [class]; |
|
722 |
fun subst_aterm (t as Free (v, ty)) = (case AList.lookup (op =) consts v |
|
723 |
of SOME (c, _) => Const (c, ty) |
|
724 |
| NONE => t) |
|
725 |
| subst_aterm t = t; |
|
726 |
in map_types subst_typ #> Term.map_aterms subst_aterm end; |
|
727 |
||
728 |
fun add_const_in_class class ((c, rhs), syn) thy = |
|
729 |
let |
|
730 |
val prfx = (Logic.const_of_class o NameSpace.base) class; |
|
731 |
fun mk_name inject c = |
|
732 |
let |
|
733 |
val n1 = Sign.full_name thy c; |
|
734 |
val n2 = NameSpace.qualifier n1; |
|
735 |
val n3 = NameSpace.base n1; |
|
736 |
in NameSpace.implode (n2 :: inject @ [n3]) end; |
|
737 |
val abbr' = mk_name [prfx, prfx] c; |
|
738 |
val rhs' = export_fixes thy class rhs; |
|
739 |
val ty' = Term.fastype_of rhs'; |
|
740 |
val def = (c, Logic.mk_equals (Const (mk_name [prfx] c, ty'), rhs')); |
|
741 |
val (syn', _) = fork_mixfix true NONE syn; |
|
742 |
fun interpret def = |
|
743 |
let |
|
744 |
val def' = symmetric def |
|
745 |
val tac = (ALLGOALS o ProofContext.fact_tac) [def']; |
|
746 |
val name_locale = (#locale o fst o the_class_data thy) class; |
|
747 |
val def_eq = Thm.prop_of def'; |
|
748 |
val (params, consts) = split_list (param_map thy [class]); |
|
749 |
in |
|
750 |
prove_interpretation tac ((false, prfx), []) (Locale.Locale name_locale) |
|
751 |
((mk_instT class, mk_inst class params consts), [def_eq]) |
|
752 |
#> add_class_const_thm (class, def') |
|
753 |
end; |
|
754 |
in |
|
755 |
thy |
|
756 |
|> Sign.hide_consts_i true [abbr'] |
|
757 |
|> Sign.add_path prfx |
|
758 |
|> Sign.add_consts_authentic [(c, ty', syn')] |
|
759 |
|> Sign.parent_path |
|
760 |
|> Sign.sticky_prefix prfx |
|
761 |
|> PureThy.add_defs_i false [(def, [])] |
|
762 |
|-> (fn [def] => interpret def) |
|
763 |
|> Sign.restore_naming thy |
|
764 |
end; |
|
765 |
||
766 |
end; |