author | paulson |
Mon, 30 May 2011 16:15:37 +0100 | |
changeset 43078 | e2631aaf1e1e |
parent 42402 | c7139609b67d |
child 43329 | 84472e198515 |
permissions | -rw-r--r-- |
38379
67d71449e85b
more convenient split of class modules: class and class_declaration
haftmann
parents:
38377
diff
changeset
|
1 |
(* Title: Pure/Isar/class.ML |
24218 | 2 |
Author: Florian Haftmann, TU Muenchen |
3 |
||
38379
67d71449e85b
more convenient split of class modules: class and class_declaration
haftmann
parents:
38377
diff
changeset
|
4 |
Type classes derived from primitive axclasses and locales. |
24218 | 5 |
*) |
6 |
||
38379
67d71449e85b
more convenient split of class modules: class and class_declaration
haftmann
parents:
38377
diff
changeset
|
7 |
signature CLASS = |
24218 | 8 |
sig |
25462 | 9 |
(*classes*) |
29526 | 10 |
val is_class: theory -> class -> bool |
11 |
val these_params: theory -> sort -> (string * (class * (string * typ))) list |
|
38379
67d71449e85b
more convenient split of class modules: class and class_declaration
haftmann
parents:
38377
diff
changeset
|
12 |
val base_sort: theory -> class -> sort |
67d71449e85b
more convenient split of class modules: class and class_declaration
haftmann
parents:
38377
diff
changeset
|
13 |
val rules: theory -> class -> thm option * thm |
67d71449e85b
more convenient split of class modules: class and class_declaration
haftmann
parents:
38377
diff
changeset
|
14 |
val these_defs: theory -> sort -> thm list |
67d71449e85b
more convenient split of class modules: class and class_declaration
haftmann
parents:
38377
diff
changeset
|
15 |
val these_operations: theory -> sort |
67d71449e85b
more convenient split of class modules: class and class_declaration
haftmann
parents:
38377
diff
changeset
|
16 |
-> (string * (class * (typ * term))) list |
42359 | 17 |
val print_classes: Proof.context -> unit |
25311 | 18 |
val init: class -> theory -> Proof.context |
38379
67d71449e85b
more convenient split of class modules: class and class_declaration
haftmann
parents:
38377
diff
changeset
|
19 |
val begin: class list -> sort -> Proof.context -> Proof.context |
38619
25e401d53900
tuned: less formal noise in Named_Target, more coherence in Class
haftmann
parents:
38392
diff
changeset
|
20 |
val const: class -> (binding * mixfix) * (term list * term) -> local_theory -> local_theory |
25e401d53900
tuned: less formal noise in Named_Target, more coherence in Class
haftmann
parents:
38392
diff
changeset
|
21 |
val abbrev: class -> Syntax.mode -> (binding * mixfix) * term -> local_theory -> local_theory |
38379
67d71449e85b
more convenient split of class modules: class and class_declaration
haftmann
parents:
38377
diff
changeset
|
22 |
val redeclare_operations: theory -> sort -> Proof.context -> Proof.context |
29526 | 23 |
val class_prefix: string -> string |
38379
67d71449e85b
more convenient split of class modules: class and class_declaration
haftmann
parents:
38377
diff
changeset
|
24 |
val register: class -> class list -> ((string * typ) * (string * typ)) list |
67d71449e85b
more convenient split of class modules: class and class_declaration
haftmann
parents:
38377
diff
changeset
|
25 |
-> sort -> morphism -> morphism -> thm option -> thm option -> thm |
67d71449e85b
more convenient split of class modules: class and class_declaration
haftmann
parents:
38377
diff
changeset
|
26 |
-> theory -> theory |
25485 | 27 |
|
25462 | 28 |
(*instances*) |
38379
67d71449e85b
more convenient split of class modules: class and class_declaration
haftmann
parents:
38377
diff
changeset
|
29 |
val instantiation: string list * (string * sort) list * sort -> theory -> local_theory |
26247 | 30 |
val instantiation_instance: (local_theory -> local_theory) |
31 |
-> local_theory -> Proof.state |
|
32 |
val prove_instantiation_instance: (Proof.context -> tactic) |
|
33 |
-> local_theory -> local_theory |
|
28666 | 34 |
val prove_instantiation_exit: (Proof.context -> tactic) |
35 |
-> local_theory -> theory |
|
36 |
val prove_instantiation_exit_result: (morphism -> 'a -> 'b) |
|
37 |
-> (Proof.context -> 'b -> tactic) -> 'a -> local_theory -> 'b * theory |
|
31869 | 38 |
val read_multi_arity: theory -> xstring list * xstring list * xstring |
39 |
-> string list * (string * sort) list * sort |
|
26259 | 40 |
val type_name: string -> string |
38348
cf7b2121ad9d
moved instantiation target formally to class_target.ML
haftmann
parents:
38107
diff
changeset
|
41 |
val instantiation_cmd: xstring list * xstring list * xstring -> theory -> local_theory |
38377 | 42 |
val instance_arity_cmd: xstring list * xstring list * xstring -> theory -> Proof.state |
25485 | 43 |
|
31635 | 44 |
(*subclasses*) |
45 |
val classrel: class * class -> theory -> Proof.state |
|
46 |
val classrel_cmd: xstring * xstring -> theory -> Proof.state |
|
38379
67d71449e85b
more convenient split of class modules: class and class_declaration
haftmann
parents:
38377
diff
changeset
|
47 |
val register_subclass: class * class -> morphism option -> Element.witness option |
67d71449e85b
more convenient split of class modules: class and class_declaration
haftmann
parents:
38377
diff
changeset
|
48 |
-> morphism -> theory -> theory |
31635 | 49 |
|
50 |
(*tactics*) |
|
29526 | 51 |
val intro_classes_tac: thm list -> tactic |
52 |
val default_intro_tac: Proof.context -> thm list -> tactic |
|
24218 | 53 |
end; |
54 |
||
38379
67d71449e85b
more convenient split of class modules: class and class_declaration
haftmann
parents:
38377
diff
changeset
|
55 |
structure Class: CLASS = |
24218 | 56 |
struct |
57 |
||
24589 | 58 |
(** class data **) |
24218 | 59 |
|
60 |
datatype class_data = ClassData of { |
|
28715
238f9966c80e
class morphism stemming from locale interpretation
haftmann
parents:
28674
diff
changeset
|
61 |
|
238f9966c80e
class morphism stemming from locale interpretation
haftmann
parents:
28674
diff
changeset
|
62 |
(* static part *) |
24218 | 63 |
consts: (string * string) list |
24836 | 64 |
(*locale parameter ~> constant name*), |
25062 | 65 |
base_sort: sort, |
29545 | 66 |
base_morph: morphism |
29439 | 67 |
(*static part of canonical morphism*), |
32850 | 68 |
export_morph: morphism, |
25618 | 69 |
assm_intro: thm option, |
70 |
of_class: thm, |
|
71 |
axiom: thm option, |
|
42359 | 72 |
|
28715
238f9966c80e
class morphism stemming from locale interpretation
haftmann
parents:
28674
diff
changeset
|
73 |
(* dynamic part *) |
24657 | 74 |
defs: thm list, |
25368
f12613fda79d
proper implementation of check phase; non-qualified names for class operations
haftmann
parents:
25344
diff
changeset
|
75 |
operations: (string * (class * (typ * term))) list |
28715
238f9966c80e
class morphism stemming from locale interpretation
haftmann
parents:
28674
diff
changeset
|
76 |
|
24657 | 77 |
}; |
24218 | 78 |
|
32850 | 79 |
fun make_class_data ((consts, base_sort, base_morph, export_morph, assm_intro, of_class, axiom), |
25368
f12613fda79d
proper implementation of check phase; non-qualified names for class operations
haftmann
parents:
25344
diff
changeset
|
80 |
(defs, operations)) = |
29526 | 81 |
ClassData { consts = consts, base_sort = base_sort, |
32850 | 82 |
base_morph = base_morph, export_morph = export_morph, assm_intro = assm_intro, |
83 |
of_class = of_class, axiom = axiom, defs = defs, operations = operations }; |
|
84 |
fun map_class_data f (ClassData { consts, base_sort, base_morph, export_morph, assm_intro, |
|
28715
238f9966c80e
class morphism stemming from locale interpretation
haftmann
parents:
28674
diff
changeset
|
85 |
of_class, axiom, defs, operations }) = |
32850 | 86 |
make_class_data (f ((consts, base_sort, base_morph, export_morph, assm_intro, of_class, axiom), |
25368
f12613fda79d
proper implementation of check phase; non-qualified names for class operations
haftmann
parents:
25344
diff
changeset
|
87 |
(defs, operations))); |
25038 | 88 |
fun merge_class_data _ (ClassData { consts = consts, |
32850 | 89 |
base_sort = base_sort, base_morph = base_morph, export_morph = export_morph, assm_intro = assm_intro, |
25618 | 90 |
of_class = of_class, axiom = axiom, defs = defs1, operations = operations1 }, |
32850 | 91 |
ClassData { consts = _, base_sort = _, base_morph = _, export_morph = _, assm_intro = _, |
25618 | 92 |
of_class = _, axiom = _, defs = defs2, operations = operations2 }) = |
32850 | 93 |
make_class_data ((consts, base_sort, base_morph, export_morph, assm_intro, of_class, axiom), |
24914
95cda5dd58d5
added proper subclass concept; improved class target
haftmann
parents:
24901
diff
changeset
|
94 |
(Thm.merge_thms (defs1, defs2), |
25368
f12613fda79d
proper implementation of check phase; non-qualified names for class operations
haftmann
parents:
25344
diff
changeset
|
95 |
AList.merge (op =) (K true) (operations1, operations2))); |
24218 | 96 |
|
33522 | 97 |
structure ClassData = Theory_Data |
24218 | 98 |
( |
25038 | 99 |
type T = class_data Graph.T |
100 |
val empty = Graph.empty; |
|
24218 | 101 |
val extend = I; |
33522 | 102 |
val merge = Graph.join merge_class_data; |
24218 | 103 |
); |
104 |
||
105 |
||
106 |
(* queries *) |
|
107 |
||
31599 | 108 |
fun lookup_class_data thy class = case try (Graph.get_node (ClassData.get thy)) class |
109 |
of SOME (ClassData data) => SOME data |
|
110 |
| NONE => NONE; |
|
24218 | 111 |
|
24589 | 112 |
fun the_class_data thy class = case lookup_class_data thy class |
25020 | 113 |
of NONE => error ("Undeclared class " ^ quote class) |
24589 | 114 |
| SOME data => data; |
24218 | 115 |
|
25038 | 116 |
val is_class = is_some oo lookup_class_data; |
117 |
||
118 |
val ancestry = Graph.all_succs o ClassData.get; |
|
29509
1ff0f3f08a7b
migrated class package to new locale implementation
haftmann
parents:
29439
diff
changeset
|
119 |
val heritage = Graph.all_preds o ClassData.get; |
1ff0f3f08a7b
migrated class package to new locale implementation
haftmann
parents:
29439
diff
changeset
|
120 |
|
25002 | 121 |
fun these_params thy = |
24218 | 122 |
let |
123 |
fun params class = |
|
124 |
let |
|
24930
cc2e0e8c81af
renamed AxClass.get_definition to AxClass.get_info (again);
wenzelm
parents:
24920
diff
changeset
|
125 |
val const_typs = (#params o AxClass.get_info thy) class; |
24657 | 126 |
val const_names = (#consts o the_class_data thy) class; |
24218 | 127 |
in |
26518 | 128 |
(map o apsnd) |
129 |
(fn c => (class, (c, (the o AList.lookup (op =) const_typs) c))) const_names |
|
24218 | 130 |
end; |
131 |
in maps params o ancestry thy end; |
|
132 |
||
29526 | 133 |
val base_sort = #base_sort oo the_class_data; |
134 |
||
135 |
fun rules thy class = |
|
136 |
let val { axiom, of_class, ... } = the_class_data thy class |
|
137 |
in (axiom, of_class) end; |
|
138 |
||
139 |
fun all_assm_intros thy = |
|
31599 | 140 |
Graph.fold (fn (_, (ClassData { assm_intro, ... }, _)) => fold (insert Thm.eq_thm) |
141 |
(the_list assm_intro)) (ClassData.get thy) []; |
|
24218 | 142 |
|
29526 | 143 |
fun these_defs thy = maps (#defs o the_class_data thy) o ancestry thy; |
144 |
fun these_operations thy = maps (#operations o the_class_data thy) o ancestry thy; |
|
29358 | 145 |
|
29526 | 146 |
val base_morphism = #base_morph oo the_class_data; |
36674
d95f39448121
eq_morphism is always optional: avoid trivial morphism for empty list of equations
haftmann
parents:
36672
diff
changeset
|
147 |
fun morphism thy class = case Element.eq_morphism thy (these_defs thy [class]) |
d95f39448121
eq_morphism is always optional: avoid trivial morphism for empty list of equations
haftmann
parents:
36672
diff
changeset
|
148 |
of SOME eq_morph => base_morphism thy class $> eq_morph |
d95f39448121
eq_morphism is always optional: avoid trivial morphism for empty list of equations
haftmann
parents:
36672
diff
changeset
|
149 |
| NONE => base_morphism thy class; |
32850 | 150 |
val export_morphism = #export_morph oo the_class_data; |
28715
238f9966c80e
class morphism stemming from locale interpretation
haftmann
parents:
28674
diff
changeset
|
151 |
|
42359 | 152 |
fun print_classes ctxt = |
24218 | 153 |
let |
42360 | 154 |
val thy = Proof_Context.theory_of ctxt; |
24218 | 155 |
val algebra = Sign.classes_of thy; |
156 |
val arities = |
|
157 |
Symtab.empty |
|
158 |
|> Symtab.fold (fn (tyco, arities) => fold (fn (class, _) => |
|
159 |
Symtab.map_default (class, []) (insert (op =) tyco)) arities) |
|
36328
4d9deabf6474
replaced Sorts.rep_algebra by slightly more abstract selectors classes_of/arities_of;
wenzelm
parents:
36323
diff
changeset
|
160 |
(Sorts.arities_of algebra); |
24218 | 161 |
val the_arities = these o Symtab.lookup arities; |
162 |
fun mk_arity class tyco = |
|
163 |
let |
|
164 |
val Ss = Sorts.mg_domain algebra tyco [class]; |
|
24920 | 165 |
in Syntax.pretty_arity ctxt (tyco, Ss, [class]) end; |
42359 | 166 |
fun mk_param (c, ty) = |
42360 | 167 |
Pretty.str (Proof_Context.extern_const ctxt c ^ " :: " ^ |
42359 | 168 |
Syntax.string_of_typ (Config.put show_sorts false ctxt) (Type.strip_sorts ty)); |
24218 | 169 |
fun mk_entry class = (Pretty.block o Pretty.fbreaks o map_filter I) [ |
42360 | 170 |
(SOME o Pretty.str) ("class " ^ Proof_Context.extern_class ctxt class ^ ":"), |
24218 | 171 |
(SOME o Pretty.block) [Pretty.str "supersort: ", |
24920 | 172 |
(Syntax.pretty_sort ctxt o Sign.minimize_sort thy o Sign.super_classes thy) class], |
25062 | 173 |
((fn [] => NONE | ps => (SOME o Pretty.block o Pretty.fbreaks) |
174 |
(Pretty.str "parameters:" :: ps)) o map mk_param |
|
24930
cc2e0e8c81af
renamed AxClass.get_definition to AxClass.get_info (again);
wenzelm
parents:
24920
diff
changeset
|
175 |
o these o Option.map #params o try (AxClass.get_info thy)) class, |
24218 | 176 |
(SOME o Pretty.block o Pretty.breaks) [ |
177 |
Pretty.str "instances:", |
|
178 |
Pretty.list "" "" (map (mk_arity class) (the_arities class)) |
|
179 |
] |
|
180 |
] |
|
181 |
in |
|
24589 | 182 |
(Pretty.writeln o Pretty.chunks o separate (Pretty.str "") |
183 |
o map mk_entry o Sorts.all_classes) algebra |
|
24218 | 184 |
end; |
185 |
||
186 |
||
187 |
(* updaters *) |
|
188 |
||
32850 | 189 |
fun register class sups params base_sort base_morph export_morph |
29358 | 190 |
axiom assm_intro of_class thy = |
25002 | 191 |
let |
25368
f12613fda79d
proper implementation of check phase; non-qualified names for class operations
haftmann
parents:
25344
diff
changeset
|
192 |
val operations = map (fn (v_ty as (_, ty), (c, _)) => |
25683 | 193 |
(c, (class, (ty, Free v_ty)))) params; |
25038 | 194 |
val add_class = Graph.new_node (class, |
31599 | 195 |
make_class_data (((map o pairself) fst params, base_sort, |
32850 | 196 |
base_morph, export_morph, assm_intro, of_class, axiom), ([], operations))) |
29526 | 197 |
#> fold (curry Graph.add_edge class) sups; |
25618 | 198 |
in ClassData.map add_class thy end; |
24218 | 199 |
|
36674
d95f39448121
eq_morphism is always optional: avoid trivial morphism for empty list of equations
haftmann
parents:
36672
diff
changeset
|
200 |
fun activate_defs class thms thy = case Element.eq_morphism thy thms |
d95f39448121
eq_morphism is always optional: avoid trivial morphism for empty list of equations
haftmann
parents:
36672
diff
changeset
|
201 |
of SOME eq_morph => fold (fn cls => fn thy => |
38107 | 202 |
Context.theory_map (Locale.amend_registration (cls, base_morphism thy cls) |
203 |
(eq_morph, true) (export_morphism thy cls)) thy) (heritage thy [class]) thy |
|
36674
d95f39448121
eq_morphism is always optional: avoid trivial morphism for empty list of equations
haftmann
parents:
36672
diff
changeset
|
204 |
| NONE => thy; |
29526 | 205 |
|
25368
f12613fda79d
proper implementation of check phase; non-qualified names for class operations
haftmann
parents:
25344
diff
changeset
|
206 |
fun register_operation class (c, (t, some_def)) thy = |
25062 | 207 |
let |
29358 | 208 |
val base_sort = base_sort thy class; |
29439 | 209 |
val prep_typ = map_type_tfree |
210 |
(fn (v, sort) => if Name.aT = v |
|
211 |
then TFree (v, base_sort) else TVar ((v, 0), sort)); |
|
25368
f12613fda79d
proper implementation of check phase; non-qualified names for class operations
haftmann
parents:
25344
diff
changeset
|
212 |
val t' = map_types prep_typ t; |
f12613fda79d
proper implementation of check phase; non-qualified names for class operations
haftmann
parents:
25344
diff
changeset
|
213 |
val ty' = Term.fastype_of t'; |
25062 | 214 |
in |
215 |
thy |
|
216 |
|> (ClassData.map o Graph.map_node class o map_class_data o apsnd) |
|
25368
f12613fda79d
proper implementation of check phase; non-qualified names for class operations
haftmann
parents:
25344
diff
changeset
|
217 |
(fn (defs, operations) => |
25096 | 218 |
(fold cons (the_list some_def) defs, |
25368
f12613fda79d
proper implementation of check phase; non-qualified names for class operations
haftmann
parents:
25344
diff
changeset
|
219 |
(c, (class, (ty', t'))) :: operations)) |
36674
d95f39448121
eq_morphism is always optional: avoid trivial morphism for empty list of equations
haftmann
parents:
36672
diff
changeset
|
220 |
|> activate_defs class (the_list some_def) |
25062 | 221 |
end; |
24218 | 222 |
|
29558 | 223 |
fun register_subclass (sub, sup) some_dep_morph some_wit export thy = |
25618 | 224 |
let |
29558 | 225 |
val intros = (snd o rules thy) sup :: map_filter I |
35021
c839a4c670c6
renamed old-style Drule.standard to Drule.export_without_context, to emphasize that this is in no way a standard operation;
wenzelm
parents:
33969
diff
changeset
|
226 |
[Option.map (Drule.export_without_context_open o Element.conclude_witness) some_wit, |
29558 | 227 |
(fst o rules thy) sub]; |
29610 | 228 |
val tac = EVERY (map (TRYALL o Tactic.rtac) intros); |
32970
fbd2bb2489a8
operations of structure Skip_Proof (formerly SkipProof) no longer require quick_and_dirty mode;
wenzelm
parents:
32966
diff
changeset
|
229 |
val classrel = Skip_Proof.prove_global thy [] [] (Logic.mk_classrel (sub, sup)) |
29558 | 230 |
(K tac); |
29509
1ff0f3f08a7b
migrated class package to new locale implementation
haftmann
parents:
29439
diff
changeset
|
231 |
val diff_sort = Sign.complete_sort thy [sup] |
31012 | 232 |
|> subtract (op =) (Sign.complete_sort thy [sub]) |
233 |
|> filter (is_class thy); |
|
32074 | 234 |
val add_dependency = case some_dep_morph |
235 |
of SOME dep_morph => Locale.add_dependency sub |
|
41270 | 236 |
(sup, dep_morph $> Element.satisfy_morphism (the_list some_wit)) NONE export |
32074 | 237 |
| NONE => I |
25618 | 238 |
in |
239 |
thy |
|
37246 | 240 |
|> AxClass.add_classrel classrel |
25618 | 241 |
|> ClassData.map (Graph.add_edge (sub, sup)) |
29526 | 242 |
|> activate_defs sub (these_defs thy diff_sort) |
32074 | 243 |
|> add_dependency |
24218 | 244 |
end; |
245 |
||
246 |
||
24589 | 247 |
(** classes and class target **) |
24218 | 248 |
|
25002 | 249 |
(* class context syntax *) |
24748 | 250 |
|
35858
0d394a82337e
handle hidden polymorphism in class target (without class target syntax, though)
haftmann
parents:
35845
diff
changeset
|
251 |
fun these_unchecks thy = map (fn (c, (_, (ty, t))) => (t, Const (c, ty))) |
0d394a82337e
handle hidden polymorphism in class target (without class target syntax, though)
haftmann
parents:
35845
diff
changeset
|
252 |
o these_operations thy; |
29577 | 253 |
|
29632
c3d576157244
fixed reading of class specs: declare class operations in context
haftmann
parents:
29610
diff
changeset
|
254 |
fun redeclare_const thy c = |
30364
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents:
30344
diff
changeset
|
255 |
let val b = Long_Name.base_name c |
29632
c3d576157244
fixed reading of class specs: declare class operations in context
haftmann
parents:
29610
diff
changeset
|
256 |
in Sign.intern_const thy b = c ? Variable.declare_const (b, c) end; |
c3d576157244
fixed reading of class specs: declare class operations in context
haftmann
parents:
29610
diff
changeset
|
257 |
|
29577 | 258 |
fun synchronize_class_syntax sort base_sort ctxt = |
24914
95cda5dd58d5
added proper subclass concept; improved class target
haftmann
parents:
24901
diff
changeset
|
259 |
let |
42360 | 260 |
val thy = Proof_Context.theory_of ctxt; |
26596 | 261 |
val algebra = Sign.classes_of thy; |
29577 | 262 |
val operations = these_operations thy sort; |
26518 | 263 |
fun subst_class_typ sort = map_type_tfree (K (TVar ((Name.aT, 0), sort))); |
264 |
val primary_constraints = |
|
25368
f12613fda79d
proper implementation of check phase; non-qualified names for class operations
haftmann
parents:
25344
diff
changeset
|
265 |
(map o apsnd) (subst_class_typ base_sort o fst o snd) operations; |
26518 | 266 |
val secondary_constraints = |
25368
f12613fda79d
proper implementation of check phase; non-qualified names for class operations
haftmann
parents:
25344
diff
changeset
|
267 |
(map o apsnd) (fn (class, (ty, _)) => subst_class_typ [class] ty) operations; |
26518 | 268 |
fun improve (c, ty) = (case AList.lookup (op =) primary_constraints c |
26238 | 269 |
of SOME ty' => (case try (Type.raw_match (ty', ty)) Vartab.empty |
270 |
of SOME tyenv => (case Vartab.lookup tyenv (Name.aT, 0) |
|
33969 | 271 |
of SOME (_, ty' as TVar (vi, sort)) => |
37145
01aa36932739
renamed structure TypeInfer to Type_Infer, keeping the old name as legacy alias for some time;
wenzelm
parents:
36674
diff
changeset
|
272 |
if Type_Infer.is_param vi |
26596 | 273 |
andalso Sorts.sort_le algebra (base_sort, sort) |
274 |
then SOME (ty', TFree (Name.aT, base_sort)) |
|
275 |
else NONE |
|
26238 | 276 |
| _ => NONE) |
277 |
| NONE => NONE) |
|
278 |
| NONE => NONE) |
|
279 |
fun subst (c, ty) = Option.map snd (AList.lookup (op =) operations c); |
|
29577 | 280 |
val unchecks = these_unchecks thy sort; |
25083 | 281 |
in |
282 |
ctxt |
|
29632
c3d576157244
fixed reading of class specs: declare class operations in context
haftmann
parents:
29610
diff
changeset
|
283 |
|> fold (redeclare_const thy o fst) primary_constraints |
26518 | 284 |
|> Overloading.map_improvable_syntax (K (((primary_constraints, secondary_constraints), |
26730 | 285 |
(((improve, subst), true), unchecks)), false)) |
26518 | 286 |
|> Overloading.set_primary_constraints |
25083 | 287 |
end; |
288 |
||
29632
c3d576157244
fixed reading of class specs: declare class operations in context
haftmann
parents:
29610
diff
changeset
|
289 |
fun redeclare_operations thy sort = |
c3d576157244
fixed reading of class specs: declare class operations in context
haftmann
parents:
29610
diff
changeset
|
290 |
fold (redeclare_const thy o fst) (these_operations thy sort); |
c3d576157244
fixed reading of class specs: declare class operations in context
haftmann
parents:
29610
diff
changeset
|
291 |
|
29577 | 292 |
fun begin sort base_sort ctxt = |
25083 | 293 |
ctxt |
294 |
|> Variable.declare_term |
|
295 |
(Logic.mk_type (TFree (Name.aT, base_sort))) |
|
29577 | 296 |
|> synchronize_class_syntax sort base_sort |
39378
df86b1b4ce10
more precise name for activation of improveable syntax
haftmann
parents:
39134
diff
changeset
|
297 |
|> Overloading.activate_improvable_syntax; |
24901
d3cbf79769b9
added first version of user-space type system for class target
haftmann
parents:
24847
diff
changeset
|
298 |
|
25311 | 299 |
fun init class thy = |
300 |
thy |
|
29509
1ff0f3f08a7b
migrated class package to new locale implementation
haftmann
parents:
29439
diff
changeset
|
301 |
|> Locale.init class |
29358 | 302 |
|> begin [class] (base_sort thy class); |
24914
95cda5dd58d5
added proper subclass concept; improved class target
haftmann
parents:
24901
diff
changeset
|
303 |
|
24748 | 304 |
|
27690
24738db98d34
some steps towards explicit class target for canonical interpretation
haftmann
parents:
27684
diff
changeset
|
305 |
(* class target *) |
24738db98d34
some steps towards explicit class target for canonical interpretation
haftmann
parents:
27684
diff
changeset
|
306 |
|
30364
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents:
30344
diff
changeset
|
307 |
val class_prefix = Logic.const_of_class o Long_Name.base_name; |
29526 | 308 |
|
38619
25e401d53900
tuned: less formal noise in Named_Target, more coherence in Class
haftmann
parents:
38392
diff
changeset
|
309 |
fun target_extension f class lthy = |
25e401d53900
tuned: less formal noise in Named_Target, more coherence in Class
haftmann
parents:
38392
diff
changeset
|
310 |
lthy |
25e401d53900
tuned: less formal noise in Named_Target, more coherence in Class
haftmann
parents:
38392
diff
changeset
|
311 |
|> Local_Theory.raw_theory f |
25e401d53900
tuned: less formal noise in Named_Target, more coherence in Class
haftmann
parents:
38392
diff
changeset
|
312 |
|> Local_Theory.target (synchronize_class_syntax [class] |
42360 | 313 |
(base_sort (Proof_Context.theory_of lthy) class)); |
38619
25e401d53900
tuned: less formal noise in Named_Target, more coherence in Class
haftmann
parents:
38392
diff
changeset
|
314 |
|
25e401d53900
tuned: less formal noise in Named_Target, more coherence in Class
haftmann
parents:
38392
diff
changeset
|
315 |
local |
25e401d53900
tuned: less formal noise in Named_Target, more coherence in Class
haftmann
parents:
38392
diff
changeset
|
316 |
|
25e401d53900
tuned: less formal noise in Named_Target, more coherence in Class
haftmann
parents:
38392
diff
changeset
|
317 |
fun target_const class ((c, mx), (type_params, dict)) thy = |
27690
24738db98d34
some steps towards explicit class target for canonical interpretation
haftmann
parents:
27684
diff
changeset
|
318 |
let |
29577 | 319 |
val morph = morphism thy class; |
30344
10a67c5ddddb
more uniform handling of binding in targets and derived elements;
wenzelm
parents:
30280
diff
changeset
|
320 |
val b = Morphism.binding morph c; |
10a67c5ddddb
more uniform handling of binding in targets and derived elements;
wenzelm
parents:
30280
diff
changeset
|
321 |
val b_def = Morphism.binding morph (Binding.suffix_name "_dict" b); |
29577 | 322 |
val c' = Sign.full_name thy b; |
29439 | 323 |
val dict' = Morphism.term morph dict; |
35858
0d394a82337e
handle hidden polymorphism in class target (without class target syntax, though)
haftmann
parents:
35845
diff
changeset
|
324 |
val ty' = map Term.fastype_of type_params ---> Term.fastype_of dict'; |
0d394a82337e
handle hidden polymorphism in class target (without class target syntax, though)
haftmann
parents:
35845
diff
changeset
|
325 |
val def_eq = Logic.mk_equals (list_comb (Const (c', ty'), type_params), dict') |
29577 | 326 |
|> map_types Type.strip_sorts; |
27690
24738db98d34
some steps towards explicit class target for canonical interpretation
haftmann
parents:
27684
diff
changeset
|
327 |
in |
29577 | 328 |
thy |
42375
774df7c59508
report Name_Space.declare/define, relatively to context;
wenzelm
parents:
42360
diff
changeset
|
329 |
|> Sign.declare_const_global ((b, Type.strip_sorts ty'), mx) |
29577 | 330 |
|> snd |
42375
774df7c59508
report Name_Space.declare/define, relatively to context;
wenzelm
parents:
42360
diff
changeset
|
331 |
|> Thm.add_def_global false false (b_def, def_eq) |
36106
19deea200358
Thm.add_axiom/add_def: return internal name of foundational axiom;
wenzelm
parents:
35858
diff
changeset
|
332 |
|>> apsnd Thm.varifyT_global |
39557
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents:
39438
diff
changeset
|
333 |
|-> (fn (_, def_thm) => Global_Theory.store_thm (b_def, def_thm) |
29577 | 334 |
#> snd |
35858
0d394a82337e
handle hidden polymorphism in class target (without class target syntax, though)
haftmann
parents:
35845
diff
changeset
|
335 |
#> null type_params ? register_operation class (c', (dict', SOME (Thm.symmetric def_thm)))) |
27690
24738db98d34
some steps towards explicit class target for canonical interpretation
haftmann
parents:
27684
diff
changeset
|
336 |
|> Sign.add_const_constraint (c', SOME ty') |
24738db98d34
some steps towards explicit class target for canonical interpretation
haftmann
parents:
27684
diff
changeset
|
337 |
end; |
24738db98d34
some steps towards explicit class target for canonical interpretation
haftmann
parents:
27684
diff
changeset
|
338 |
|
38619
25e401d53900
tuned: less formal noise in Named_Target, more coherence in Class
haftmann
parents:
38392
diff
changeset
|
339 |
fun target_abbrev class prmode ((c, mx), rhs) thy = |
27690
24738db98d34
some steps towards explicit class target for canonical interpretation
haftmann
parents:
27684
diff
changeset
|
340 |
let |
29577 | 341 |
val morph = morphism thy class; |
342 |
val unchecks = these_unchecks thy [class]; |
|
30344
10a67c5ddddb
more uniform handling of binding in targets and derived elements;
wenzelm
parents:
30280
diff
changeset
|
343 |
val b = Morphism.binding morph c; |
29577 | 344 |
val c' = Sign.full_name thy b; |
27690
24738db98d34
some steps towards explicit class target for canonical interpretation
haftmann
parents:
27684
diff
changeset
|
345 |
val rhs' = Pattern.rewrite_term thy unchecks [] rhs; |
24738db98d34
some steps towards explicit class target for canonical interpretation
haftmann
parents:
27684
diff
changeset
|
346 |
val ty' = Term.fastype_of rhs'; |
35845
e5980f0ad025
renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
wenzelm
parents:
35669
diff
changeset
|
347 |
val rhs'' = map_types Logic.varifyT_global rhs'; |
27690
24738db98d34
some steps towards explicit class target for canonical interpretation
haftmann
parents:
27684
diff
changeset
|
348 |
in |
29577 | 349 |
thy |
33173
b8ca12f6681a
eliminated obsolete tags for types/consts -- now handled via name space, in strongly typed fashion;
wenzelm
parents:
32970
diff
changeset
|
350 |
|> Sign.add_abbrev (#1 prmode) (b, rhs'') |
29577 | 351 |
|> snd |
27690
24738db98d34
some steps towards explicit class target for canonical interpretation
haftmann
parents:
27684
diff
changeset
|
352 |
|> Sign.add_const_constraint (c', SOME ty') |
24738db98d34
some steps towards explicit class target for canonical interpretation
haftmann
parents:
27684
diff
changeset
|
353 |
|> Sign.notation true prmode [(Const (c', ty'), mx)] |
37146
f652333bbf8e
renamed structure PrintMode to Print_Mode, keeping the old name as legacy alias for some time;
wenzelm
parents:
37145
diff
changeset
|
354 |
|> not (#1 prmode = Print_Mode.input) ? register_operation class (c', (rhs', NONE)) |
27690
24738db98d34
some steps towards explicit class target for canonical interpretation
haftmann
parents:
27684
diff
changeset
|
355 |
end; |
24738db98d34
some steps towards explicit class target for canonical interpretation
haftmann
parents:
27684
diff
changeset
|
356 |
|
38619
25e401d53900
tuned: less formal noise in Named_Target, more coherence in Class
haftmann
parents:
38392
diff
changeset
|
357 |
in |
25e401d53900
tuned: less formal noise in Named_Target, more coherence in Class
haftmann
parents:
38392
diff
changeset
|
358 |
|
25e401d53900
tuned: less formal noise in Named_Target, more coherence in Class
haftmann
parents:
38392
diff
changeset
|
359 |
fun const class arg = target_extension (target_const class arg) class; |
25e401d53900
tuned: less formal noise in Named_Target, more coherence in Class
haftmann
parents:
38392
diff
changeset
|
360 |
fun abbrev class prmode arg = target_extension (target_abbrev class prmode arg) class; |
25e401d53900
tuned: less formal noise in Named_Target, more coherence in Class
haftmann
parents:
38392
diff
changeset
|
361 |
|
25e401d53900
tuned: less formal noise in Named_Target, more coherence in Class
haftmann
parents:
38392
diff
changeset
|
362 |
end; |
25e401d53900
tuned: less formal noise in Named_Target, more coherence in Class
haftmann
parents:
38392
diff
changeset
|
363 |
|
27690
24738db98d34
some steps towards explicit class target for canonical interpretation
haftmann
parents:
27684
diff
changeset
|
364 |
|
31635 | 365 |
(* simple subclasses *) |
366 |
||
367 |
local |
|
368 |
||
369 |
fun gen_classrel mk_prop classrel thy = |
|
370 |
let |
|
371 |
fun after_qed results = |
|
42360 | 372 |
Proof_Context.background_theory ((fold o fold) AxClass.add_classrel results); |
31635 | 373 |
in |
374 |
thy |
|
42360 | 375 |
|> Proof_Context.init_global |
36323
655e2d74de3a
modernized naming conventions of main Isar proof elements;
wenzelm
parents:
36106
diff
changeset
|
376 |
|> Proof.theorem NONE after_qed [[(mk_prop thy classrel, [])]] |
31635 | 377 |
end; |
378 |
||
379 |
in |
|
380 |
||
381 |
val classrel = |
|
382 |
gen_classrel (Logic.mk_classrel oo AxClass.cert_classrel); |
|
383 |
val classrel_cmd = |
|
384 |
gen_classrel (Logic.mk_classrel oo AxClass.read_classrel); |
|
385 |
||
386 |
end; (*local*) |
|
387 |
||
388 |
||
25462 | 389 |
(** instantiation target **) |
390 |
||
391 |
(* bookkeeping *) |
|
392 |
||
393 |
datatype instantiation = Instantiation of { |
|
25864 | 394 |
arities: string list * (string * sort) list * sort, |
25462 | 395 |
params: ((string * string) * (string * typ)) list |
25603 | 396 |
(*(instantiation parameter, type constructor), (local instantiation parameter, typ)*) |
25462 | 397 |
} |
398 |
||
33519 | 399 |
structure Instantiation = Proof_Data |
25462 | 400 |
( |
401 |
type T = instantiation |
|
25536 | 402 |
fun init _ = Instantiation { arities = ([], [], []), params = [] }; |
25462 | 403 |
); |
404 |
||
25485 | 405 |
fun mk_instantiation (arities, params) = |
406 |
Instantiation { arities = arities, params = params }; |
|
33671 | 407 |
fun get_instantiation lthy = case Instantiation.get (Local_Theory.target_of lthy) |
25485 | 408 |
of Instantiation data => data; |
33671 | 409 |
fun map_instantiation f = (Local_Theory.target o Instantiation.map) |
25514 | 410 |
(fn Instantiation { arities, params } => mk_instantiation (f (arities, params))); |
25462 | 411 |
|
25514 | 412 |
fun the_instantiation lthy = case get_instantiation lthy |
25536 | 413 |
of { arities = ([], [], []), ... } => error "No instantiation target" |
25485 | 414 |
| data => data; |
25462 | 415 |
|
25485 | 416 |
val instantiation_params = #params o get_instantiation; |
25462 | 417 |
|
30519
c05c0199826f
coherent binding policy with primitive target operations
haftmann
parents:
30364
diff
changeset
|
418 |
fun instantiation_param lthy b = instantiation_params lthy |
c05c0199826f
coherent binding policy with primitive target operations
haftmann
parents:
30364
diff
changeset
|
419 |
|> find_first (fn (_, (v, _)) => Binding.name_of b = v) |
25462 | 420 |
|> Option.map (fst o fst); |
421 |
||
31869 | 422 |
fun read_multi_arity thy (raw_tycos, raw_sorts, raw_sort) = |
423 |
let |
|
42360 | 424 |
val ctxt = Proof_Context.init_global thy; |
425 |
val all_arities = map (fn raw_tyco => Proof_Context.read_arity ctxt |
|
31869 | 426 |
(raw_tyco, raw_sorts, raw_sort)) raw_tycos; |
427 |
val tycos = map #1 all_arities; |
|
428 |
val (_, sorts, sort) = hd all_arities; |
|
429 |
val vs = Name.names Name.context Name.aT sorts; |
|
430 |
in (tycos, vs, sort) end; |
|
431 |
||
25462 | 432 |
|
433 |
(* syntax *) |
|
434 |
||
26238 | 435 |
fun synchronize_inst_syntax ctxt = |
25462 | 436 |
let |
33969 | 437 |
val Instantiation { params, ... } = Instantiation.get ctxt; |
31249 | 438 |
|
33969 | 439 |
val lookup_inst_param = AxClass.lookup_inst_param |
42360 | 440 |
(Sign.consts_of (Proof_Context.theory_of ctxt)) params; |
31249 | 441 |
fun subst (c, ty) = case lookup_inst_param (c, ty) |
442 |
of SOME (v_ty as (_, ty)) => SOME (ty, Free v_ty) |
|
443 |
| NONE => NONE; |
|
26238 | 444 |
val unchecks = |
445 |
map (fn ((c, _), v_ty as (_, ty)) => (Free v_ty, Const (c, ty))) params; |
|
446 |
in |
|
447 |
ctxt |
|
26259 | 448 |
|> Overloading.map_improvable_syntax |
26730 | 449 |
(fn (((primary_constraints, _), (((improve, _), _), _)), _) => |
450 |
(((primary_constraints, []), (((improve, subst), false), unchecks)), false)) |
|
26238 | 451 |
end; |
25462 | 452 |
|
42385
b46b47775cbe
simplified Sorts.class_error: plain Proof.context;
wenzelm
parents:
42383
diff
changeset
|
453 |
fun resort_terms ctxt algebra consts constraints ts = |
38382 | 454 |
let |
42385
b46b47775cbe
simplified Sorts.class_error: plain Proof.context;
wenzelm
parents:
42383
diff
changeset
|
455 |
fun matchings (Const (c_ty as (c, _))) = |
b46b47775cbe
simplified Sorts.class_error: plain Proof.context;
wenzelm
parents:
42383
diff
changeset
|
456 |
(case constraints c of |
b46b47775cbe
simplified Sorts.class_error: plain Proof.context;
wenzelm
parents:
42383
diff
changeset
|
457 |
NONE => I |
b46b47775cbe
simplified Sorts.class_error: plain Proof.context;
wenzelm
parents:
42383
diff
changeset
|
458 |
| SOME sorts => |
b46b47775cbe
simplified Sorts.class_error: plain Proof.context;
wenzelm
parents:
42383
diff
changeset
|
459 |
fold2 (curry (Sorts.meet_sort algebra)) (Consts.typargs consts c_ty) sorts) |
b46b47775cbe
simplified Sorts.class_error: plain Proof.context;
wenzelm
parents:
42383
diff
changeset
|
460 |
| matchings _ = I; |
38382 | 461 |
val tvartab = (fold o fold_aterms) matchings ts Vartab.empty |
42385
b46b47775cbe
simplified Sorts.class_error: plain Proof.context;
wenzelm
parents:
42383
diff
changeset
|
462 |
handle Sorts.CLASS_ERROR e => error (Sorts.class_error ctxt e); |
38382 | 463 |
val inst = map_type_tvar |
464 |
(fn (vi, sort) => TVar (vi, the_default sort (Vartab.lookup tvartab vi))); |
|
42402
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
wenzelm
parents:
42388
diff
changeset
|
465 |
in if Vartab.is_empty tvartab then ts else (map o map_types) inst ts end; |
38382 | 466 |
|
25462 | 467 |
|
468 |
(* target *) |
|
469 |
||
37678
0040bafffdef
"prod" and "sum" replace "*" and "+" respectively
haftmann
parents:
37393
diff
changeset
|
470 |
val sanitize_name = (*necessary as long as "dirty" type identifiers are permitted*) |
25485 | 471 |
let |
25574 | 472 |
fun is_valid s = Symbol.is_ascii_letter s orelse Symbol.is_ascii_digit s |
473 |
orelse s = "'" orelse s = "_"; |
|
25485 | 474 |
val is_junk = not o is_valid andf Symbol.is_regular; |
475 |
val junk = Scan.many is_junk; |
|
476 |
val scan_valids = Symbol.scanner "Malformed input" |
|
477 |
((junk |-- |
|
478 |
(Scan.optional (Scan.one Symbol.is_ascii_letter) "x" ^^ (Scan.many is_valid >> implode) |
|
479 |
--| junk)) |
|
25999 | 480 |
::: Scan.repeat ((Scan.many1 is_valid >> implode) --| junk)); |
25485 | 481 |
in |
40627
becf5d5187cc
renamed raw "explode" function to "raw_explode" to emphasize its meaning;
wenzelm
parents:
39557
diff
changeset
|
482 |
raw_explode #> scan_valids #> implode |
25485 | 483 |
end; |
484 |
||
37678
0040bafffdef
"prod" and "sum" replace "*" and "+" respectively
haftmann
parents:
37393
diff
changeset
|
485 |
val type_name = sanitize_name o Long_Name.base_name; |
26259 | 486 |
|
38757
2b3e054ae6fc
renamed Local_Theory.theory(_result) to Local_Theory.background_theory(_result) to emphasize that this belongs to the infrastructure and is rarely appropriate in user-space tools;
wenzelm
parents:
38756
diff
changeset
|
487 |
fun define_overloaded (c, U) v (b_def, rhs) = Local_Theory.background_theory_result |
38382 | 488 |
(AxClass.declare_overloaded (c, U) ##>> AxClass.define_overloaded b_def (c, rhs)) |
489 |
##> (map_instantiation o apsnd) (filter_out (fn (_, (v', _)) => v' = v)) |
|
490 |
##> Local_Theory.target synchronize_inst_syntax; |
|
491 |
||
492 |
fun foundation (((b, U), mx), (b_def, rhs)) (type_params, term_params) lthy = |
|
493 |
case instantiation_param lthy b |
|
494 |
of SOME c => if mx <> NoSyn then error ("Illegal mixfix syntax for overloaded constant " ^ quote c) |
|
495 |
else lthy |> define_overloaded (c, U) (Binding.name_of b) (b_def, rhs) |
|
496 |
| NONE => lthy |> |
|
497 |
Generic_Target.theory_foundation (((b, U), mx), (b_def, rhs)) (type_params, term_params); |
|
498 |
||
499 |
fun pretty lthy = |
|
26518 | 500 |
let |
38382 | 501 |
val { arities = (tycos, vs, sort), params } = the_instantiation lthy; |
42360 | 502 |
val thy = Proof_Context.theory_of lthy; |
38382 | 503 |
fun pr_arity tyco = Syntax.pretty_arity lthy (tyco, map snd vs, sort); |
504 |
fun pr_param ((c, _), (v, ty)) = |
|
42359 | 505 |
Pretty.block (Pretty.breaks |
42360 | 506 |
[Pretty.str v, Pretty.str "==", Pretty.str (Proof_Context.extern_const lthy c), |
42359 | 507 |
Pretty.str "::", Syntax.pretty_typ lthy ty]); |
38382 | 508 |
in Pretty.str "instantiation" :: map pr_arity tycos @ map pr_param params end; |
26518 | 509 |
|
38382 | 510 |
fun conclude lthy = |
511 |
let |
|
42359 | 512 |
val (tycos, vs, sort) = #arities (the_instantiation lthy); |
42360 | 513 |
val thy = Proof_Context.theory_of lthy; |
42359 | 514 |
val _ = tycos |> List.app (fn tyco => |
515 |
if Sign.of_sort thy |
|
38382 | 516 |
(Type (tyco, map TFree vs), sort) |
42359 | 517 |
then () |
42360 | 518 |
else error ("Missing instance proof for type " ^ quote (Proof_Context.extern_type lthy tyco))); |
38382 | 519 |
in lthy end; |
520 |
||
521 |
fun instantiation (tycos, vs, sort) thy = |
|
25462 | 522 |
let |
25536 | 523 |
val _ = if null tycos then error "At least one arity must be given" else (); |
31249 | 524 |
val class_params = these_params thy (filter (can (AxClass.get_info thy)) sort); |
29632
c3d576157244
fixed reading of class specs: declare class operations in context
haftmann
parents:
29610
diff
changeset
|
525 |
fun get_param tyco (param, (_, (c, ty))) = |
c3d576157244
fixed reading of class specs: declare class operations in context
haftmann
parents:
29610
diff
changeset
|
526 |
if can (AxClass.param_of_inst thy) (c, tyco) |
25603 | 527 |
then NONE else SOME ((c, tyco), |
25864 | 528 |
(param ^ "_" ^ type_name tyco, map_atyps (K (Type (tyco, map TFree vs))) ty)); |
31249 | 529 |
val params = map_product get_param tycos class_params |> map_filter I; |
26518 | 530 |
val primary_constraints = map (apsnd |
31249 | 531 |
(map_atyps (K (TVar ((Name.aT, 0), [])))) o snd o snd) class_params; |
26518 | 532 |
val algebra = Sign.classes_of thy |
42387 | 533 |
|> fold (fn tyco => Sorts.add_arities (Proof_Context.init_global thy) |
26518 | 534 |
(tyco, map (fn class => (class, map snd vs)) sort)) tycos; |
535 |
val consts = Sign.consts_of thy; |
|
536 |
val improve_constraints = AList.lookup (op =) |
|
31249 | 537 |
(map (fn (_, (class, (c, _))) => (c, [[class]])) class_params); |
42402
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
wenzelm
parents:
42388
diff
changeset
|
538 |
fun resort_check ctxt ts = resort_terms ctxt algebra consts improve_constraints ts; |
31249 | 539 |
val lookup_inst_param = AxClass.lookup_inst_param consts params; |
42388 | 540 |
fun improve (c, ty) = |
541 |
(case lookup_inst_param (c, ty) of |
|
542 |
SOME (_, ty') => if Sign.typ_instance thy (ty', ty) then SOME (ty, ty') else NONE |
|
543 |
| NONE => NONE); |
|
25485 | 544 |
in |
545 |
thy |
|
32379
a97e9caebd60
additional checkpoints avoid problems in error situations
haftmann
parents:
32074
diff
changeset
|
546 |
|> Theory.checkpoint |
42360 | 547 |
|> Proof_Context.init_global |
31249 | 548 |
|> Instantiation.put (mk_instantiation ((tycos, vs, sort), params)) |
27281 | 549 |
|> fold (Variable.declare_typ o TFree) vs |
31249 | 550 |
|> fold (Variable.declare_names o Free o snd) params |
26259 | 551 |
|> (Overloading.map_improvable_syntax o apfst) |
29969
9dbb046136d0
more precise improvement in instantiation user space type system
haftmann
parents:
29632
diff
changeset
|
552 |
(K ((primary_constraints, []), (((improve, K NONE), false), []))) |
39378
df86b1b4ce10
more precise name for activation of improveable syntax
haftmann
parents:
39134
diff
changeset
|
553 |
|> Overloading.activate_improvable_syntax |
26518 | 554 |
|> Context.proof_map (Syntax.add_term_check 0 "resorting" resort_check) |
26238 | 555 |
|> synchronize_inst_syntax |
38382 | 556 |
|> Local_Theory.init NONE "" |
557 |
{define = Generic_Target.define foundation, |
|
558 |
notes = Generic_Target.notes |
|
559 |
(fn kind => fn global_facts => fn _ => Generic_Target.theory_notes kind global_facts), |
|
560 |
abbrev = Generic_Target.abbrev |
|
42385
b46b47775cbe
simplified Sorts.class_error: plain Proof.context;
wenzelm
parents:
42383
diff
changeset
|
561 |
(fn prmode => fn (b, mx) => fn (t, _) => fn _ => |
b46b47775cbe
simplified Sorts.class_error: plain Proof.context;
wenzelm
parents:
42383
diff
changeset
|
562 |
Generic_Target.theory_abbrev prmode ((b, mx), t)), |
38382 | 563 |
declaration = K Generic_Target.theory_declaration, |
564 |
syntax_declaration = K Generic_Target.theory_declaration, |
|
565 |
pretty = pretty, |
|
566 |
exit = Local_Theory.target_of o conclude} |
|
25485 | 567 |
end; |
568 |
||
38382 | 569 |
fun instantiation_cmd arities thy = |
570 |
instantiation (read_multi_arity thy arities) thy; |
|
26238 | 571 |
|
25485 | 572 |
fun gen_instantiation_instance do_proof after_qed lthy = |
573 |
let |
|
25864 | 574 |
val (tycos, vs, sort) = (#arities o the_instantiation) lthy; |
575 |
val arities_proof = maps (fn tyco => Logic.mk_arities (tyco, map snd vs, sort)) tycos; |
|
25462 | 576 |
fun after_qed' results = |
38757
2b3e054ae6fc
renamed Local_Theory.theory(_result) to Local_Theory.background_theory(_result) to emphasize that this belongs to the infrastructure and is rarely appropriate in user-space tools;
wenzelm
parents:
38756
diff
changeset
|
577 |
Local_Theory.background_theory (fold (AxClass.add_arity o Thm.varifyT_global) results) |
25462 | 578 |
#> after_qed; |
579 |
in |
|
580 |
lthy |
|
581 |
|> do_proof after_qed' arities_proof |
|
582 |
end; |
|
583 |
||
25485 | 584 |
val instantiation_instance = gen_instantiation_instance (fn after_qed => fn ts => |
36323
655e2d74de3a
modernized naming conventions of main Isar proof elements;
wenzelm
parents:
36106
diff
changeset
|
585 |
Proof.theorem NONE (after_qed o map the_single) (map (fn t => [(t, [])]) ts)); |
25462 | 586 |
|
25485 | 587 |
fun prove_instantiation_instance tac = gen_instantiation_instance (fn after_qed => |
25502 | 588 |
fn ts => fn lthy => after_qed (map (fn t => Goal.prove lthy [] [] t |
589 |
(fn {context, ...} => tac context)) ts) lthy) I; |
|
25462 | 590 |
|
28666 | 591 |
fun prove_instantiation_exit tac = prove_instantiation_instance tac |
33671 | 592 |
#> Local_Theory.exit_global; |
28666 | 593 |
|
594 |
fun prove_instantiation_exit_result f tac x lthy = |
|
595 |
let |
|
42360 | 596 |
val morph = Proof_Context.export_morphism lthy |
597 |
(Proof_Context.init_global (Proof_Context.theory_of lthy)); |
|
29439 | 598 |
val y = f morph x; |
28666 | 599 |
in |
600 |
lthy |
|
601 |
|> prove_instantiation_exit (fn ctxt => tac ctxt y) |
|
602 |
|> pair y |
|
603 |
end; |
|
604 |
||
29526 | 605 |
|
31635 | 606 |
(* simplified instantiation interface with no class parameter *) |
607 |
||
31869 | 608 |
fun instance_arity_cmd raw_arities thy = |
31635 | 609 |
let |
31869 | 610 |
val (tycos, vs, sort) = read_multi_arity thy raw_arities; |
611 |
val sorts = map snd vs; |
|
612 |
val arities = maps (fn tyco => Logic.mk_arities (tyco, sorts, sort)) tycos; |
|
42360 | 613 |
fun after_qed results = Proof_Context.background_theory |
37246 | 614 |
((fold o fold) AxClass.add_arity results); |
31635 | 615 |
in |
616 |
thy |
|
42360 | 617 |
|> Proof_Context.init_global |
36323
655e2d74de3a
modernized naming conventions of main Isar proof elements;
wenzelm
parents:
36106
diff
changeset
|
618 |
|> Proof.theorem NONE after_qed (map (fn t => [(t, [])]) arities) |
31635 | 619 |
end; |
620 |
||
621 |
||
29526 | 622 |
(** tactics and methods **) |
623 |
||
624 |
fun intro_classes_tac facts st = |
|
625 |
let |
|
626 |
val thy = Thm.theory_of_thm st; |
|
627 |
val classes = Sign.all_classes thy; |
|
628 |
val class_trivs = map (Thm.class_triv thy) classes; |
|
629 |
val class_intros = map_filter (try (#intro o AxClass.get_info thy)) classes; |
|
630 |
val assm_intros = all_assm_intros thy; |
|
631 |
in |
|
632 |
Method.intros_tac (class_trivs @ class_intros @ assm_intros) facts st |
|
633 |
end; |
|
634 |
||
635 |
fun default_intro_tac ctxt [] = |
|
39438
c5ece2a7a86e
Isar "default" step needs to fail for solved problems, for clear distinction of '.' and '..' for example -- amending lapse introduced in 9de4d64eee3b (April 2004);
wenzelm
parents:
39378
diff
changeset
|
636 |
COND Thm.no_prems no_tac |
c5ece2a7a86e
Isar "default" step needs to fail for solved problems, for clear distinction of '.' and '..' for example -- amending lapse introduced in 9de4d64eee3b (April 2004);
wenzelm
parents:
39378
diff
changeset
|
637 |
(intro_classes_tac [] ORELSE Locale.intro_locales_tac true ctxt []) |
29526 | 638 |
| default_intro_tac _ _ = no_tac; |
639 |
||
640 |
fun default_tac rules ctxt facts = |
|
641 |
HEADGOAL (Method.some_rule_tac rules ctxt facts) ORELSE |
|
642 |
default_intro_tac ctxt facts; |
|
643 |
||
644 |
val _ = Context.>> (Context.map_theory |
|
30515 | 645 |
(Method.setup (Binding.name "intro_classes") (Scan.succeed (K (METHOD intro_classes_tac))) |
646 |
"back-chain introduction rules of classes" #> |
|
647 |
Method.setup (Binding.name "default") (Attrib.thms >> (METHOD oo default_tac)) |
|
648 |
"apply some intro/elim rule")); |
|
29526 | 649 |
|
24218 | 650 |
end; |
25683 | 651 |