author | wenzelm |
Mon, 11 Sep 2023 19:30:48 +0200 | |
changeset 78659 | b5f3d1051b13 |
parent 78453 | 3fdf3c5cfa9d |
child 79411 | 700d4f16b5f2 |
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 |
63347 | 15 |
val these_operations: theory -> sort -> (string * (class * ((typ * term) * bool))) list |
42359 | 16 |
val print_classes: Proof.context -> unit |
25311 | 17 |
val init: class -> theory -> Proof.context |
38379
67d71449e85b
more convenient split of class modules: class and class_declaration
haftmann
parents:
38377
diff
changeset
|
18 |
val begin: class list -> sort -> Proof.context -> Proof.context |
63347 | 19 |
val const: class -> (binding * mixfix) * term -> term list * term list -> |
20 |
local_theory -> local_theory |
|
21 |
val abbrev: class -> Syntax.mode -> (binding * mixfix) * term -> local_theory -> |
|
22 |
(term * term) * local_theory |
|
38379
67d71449e85b
more convenient split of class modules: class and class_declaration
haftmann
parents:
38377
diff
changeset
|
23 |
val redeclare_operations: theory -> sort -> Proof.context -> Proof.context |
29526 | 24 |
val class_prefix: string -> string |
63347 | 25 |
val register: class -> class list -> ((string * typ) * (string * typ)) list -> |
26 |
sort -> morphism -> morphism -> thm option -> thm option -> thm -> 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 |
|
38348
cf7b2121ad9d
moved instantiation target formally to class_target.ML
haftmann
parents:
38107
diff
changeset
|
40 |
val instantiation_cmd: xstring list * xstring list * xstring -> theory -> local_theory |
38377 | 41 |
val instance_arity_cmd: xstring list * xstring list * xstring -> theory -> Proof.state |
69829 | 42 |
val theory_map_result: string list * (string * sort) list * sort |
43 |
-> (morphism -> 'a -> 'b) -> (local_theory -> 'a * local_theory) |
|
44 |
-> (Proof.context -> 'b -> tactic) -> theory -> 'b * theory |
|
25485 | 45 |
|
31635 | 46 |
(*subclasses*) |
47 |
val classrel: class * class -> theory -> Proof.state |
|
48 |
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
|
49 |
val register_subclass: class * class -> morphism option -> Element.witness option |
54866
7b9a67cbd48f
self-contained formulation of subclass command, avoiding hard-wired Named_Target.init
haftmann
parents:
54742
diff
changeset
|
50 |
-> morphism -> local_theory -> local_theory |
31635 | 51 |
|
52 |
(*tactics*) |
|
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
59423
diff
changeset
|
53 |
val intro_classes_tac: Proof.context -> thm list -> tactic |
60618
4c79543cc376
renamed "default" to "standard", to make semantically clear what it is;
wenzelm
parents:
60347
diff
changeset
|
54 |
val standard_intro_classes_tac: Proof.context -> thm list -> tactic |
59296
002d817b4c37
formal pretty bodies for class specifications, accepting additional formal bookkeeping in locale.ML
haftmann
parents:
59150
diff
changeset
|
55 |
|
002d817b4c37
formal pretty bodies for class specifications, accepting additional formal bookkeeping in locale.ML
haftmann
parents:
59150
diff
changeset
|
56 |
(*diagnostics*) |
59423
3a0044e95eba
backed out obsolete workaround from ef1edfb36af7
haftmann
parents:
59420
diff
changeset
|
57 |
val pretty_specification: theory -> class -> Pretty.T list |
24218 | 58 |
end; |
59 |
||
38379
67d71449e85b
more convenient split of class modules: class and class_declaration
haftmann
parents:
38377
diff
changeset
|
60 |
structure Class: CLASS = |
24218 | 61 |
struct |
62 |
||
24589 | 63 |
(** class data **) |
24218 | 64 |
|
46919 | 65 |
datatype class_data = Class_Data of { |
28715
238f9966c80e
class morphism stemming from locale interpretation
haftmann
parents:
28674
diff
changeset
|
66 |
|
238f9966c80e
class morphism stemming from locale interpretation
haftmann
parents:
28674
diff
changeset
|
67 |
(* static part *) |
24218 | 68 |
consts: (string * string) list |
24836 | 69 |
(*locale parameter ~> constant name*), |
25062 | 70 |
base_sort: sort, |
29545 | 71 |
base_morph: morphism |
29439 | 72 |
(*static part of canonical morphism*), |
32850 | 73 |
export_morph: morphism, |
25618 | 74 |
assm_intro: thm option, |
75 |
of_class: thm, |
|
76 |
axiom: thm option, |
|
42359 | 77 |
|
28715
238f9966c80e
class morphism stemming from locale interpretation
haftmann
parents:
28674
diff
changeset
|
78 |
(* dynamic part *) |
74150 | 79 |
defs: thm Item_Net.T, |
60346 | 80 |
operations: (string * (class * ((typ * term) * bool))) list |
28715
238f9966c80e
class morphism stemming from locale interpretation
haftmann
parents:
28674
diff
changeset
|
81 |
|
48106 | 82 |
(* n.b. |
83 |
params = logical parameters of class |
|
84 |
operations = operations participating in user-space type system |
|
85 |
*) |
|
24657 | 86 |
}; |
24218 | 87 |
|
32850 | 88 |
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
|
89 |
(defs, operations)) = |
46919 | 90 |
Class_Data {consts = consts, base_sort = base_sort, |
32850 | 91 |
base_morph = base_morph, export_morph = export_morph, assm_intro = assm_intro, |
46919 | 92 |
of_class = of_class, axiom = axiom, defs = defs, operations = operations}; |
67625 | 93 |
|
46919 | 94 |
fun map_class_data f (Class_Data {consts, base_sort, base_morph, export_morph, assm_intro, |
95 |
of_class, axiom, defs, operations}) = |
|
32850 | 96 |
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
|
97 |
(defs, operations))); |
67625 | 98 |
|
46919 | 99 |
fun merge_class_data _ (Class_Data {consts = consts, |
32850 | 100 |
base_sort = base_sort, base_morph = base_morph, export_morph = export_morph, assm_intro = assm_intro, |
46919 | 101 |
of_class = of_class, axiom = axiom, defs = defs1, operations = operations1}, |
102 |
Class_Data {consts = _, base_sort = _, base_morph = _, export_morph = _, assm_intro = _, |
|
103 |
of_class = _, axiom = _, defs = defs2, operations = operations2}) = |
|
32850 | 104 |
make_class_data ((consts, base_sort, base_morph, export_morph, assm_intro, of_class, axiom), |
74150 | 105 |
(Item_Net.merge (defs1, defs2), |
25368
f12613fda79d
proper implementation of check phase; non-qualified names for class operations
haftmann
parents:
25344
diff
changeset
|
106 |
AList.merge (op =) (K true) (operations1, operations2))); |
24218 | 107 |
|
46919 | 108 |
structure Class_Data = Theory_Data |
24218 | 109 |
( |
25038 | 110 |
type T = class_data Graph.T |
111 |
val empty = Graph.empty; |
|
33522 | 112 |
val merge = Graph.join merge_class_data; |
24218 | 113 |
); |
114 |
||
115 |
||
116 |
(* queries *) |
|
117 |
||
46919 | 118 |
fun lookup_class_data thy class = |
119 |
(case try (Graph.get_node (Class_Data.get thy)) class of |
|
120 |
SOME (Class_Data data) => SOME data |
|
121 |
| NONE => NONE); |
|
24218 | 122 |
|
46919 | 123 |
fun the_class_data thy class = |
77908
a6bd716a6124
tuned: concise combinators instead of bulky case-expressions;
wenzelm
parents:
74561
diff
changeset
|
124 |
lookup_class_data thy class |
a6bd716a6124
tuned: concise combinators instead of bulky case-expressions;
wenzelm
parents:
74561
diff
changeset
|
125 |
|> \<^if_none>\<open>error ("Undeclared class " ^ quote class)\<close>; |
24218 | 126 |
|
25038 | 127 |
val is_class = is_some oo lookup_class_data; |
128 |
||
46919 | 129 |
val ancestry = Graph.all_succs o Class_Data.get; |
130 |
val heritage = Graph.all_preds o Class_Data.get; |
|
29509
1ff0f3f08a7b
migrated class package to new locale implementation
haftmann
parents:
29439
diff
changeset
|
131 |
|
25002 | 132 |
fun these_params thy = |
24218 | 133 |
let |
134 |
fun params class = |
|
135 |
let |
|
51685
385ef6706252
more standard module name Axclass (according to file name);
wenzelm
parents:
51578
diff
changeset
|
136 |
val const_typs = (#params o Axclass.get_info thy) class; |
24657 | 137 |
val const_names = (#consts o the_class_data thy) class; |
24218 | 138 |
in |
26518 | 139 |
(map o apsnd) |
140 |
(fn c => (class, (c, (the o AList.lookup (op =) const_typs) c))) const_names |
|
24218 | 141 |
end; |
142 |
in maps params o ancestry thy end; |
|
143 |
||
29526 | 144 |
val base_sort = #base_sort oo the_class_data; |
145 |
||
146 |
fun rules thy class = |
|
46919 | 147 |
let val {axiom, of_class, ...} = the_class_data thy class |
29526 | 148 |
in (axiom, of_class) end; |
149 |
||
150 |
fun all_assm_intros thy = |
|
46919 | 151 |
Graph.fold (fn (_, (Class_Data {assm_intro, ...}, _)) => fold (insert Thm.eq_thm) |
152 |
(the_list assm_intro)) (Class_Data.get thy) []; |
|
24218 | 153 |
|
74150 | 154 |
fun these_defs thy = maps (Item_Net.content o #defs o the_class_data thy) o ancestry thy; |
29526 | 155 |
fun these_operations thy = maps (#operations o the_class_data thy) o ancestry thy; |
29358 | 156 |
|
29526 | 157 |
val base_morphism = #base_morph oo the_class_data; |
47078 | 158 |
|
78453
3fdf3c5cfa9d
performance tuning: prefer static simpset within functional closure of morphism (with notable impact on specifications "in" class, e.g. AFP/No_FTL_observers);
wenzelm
parents:
78065
diff
changeset
|
159 |
fun morphism ctxt class = |
3fdf3c5cfa9d
performance tuning: prefer static simpset within functional closure of morphism (with notable impact on specifications "in" class, e.g. AFP/No_FTL_observers);
wenzelm
parents:
78065
diff
changeset
|
160 |
let val thy = Proof_Context.theory_of ctxt in |
3fdf3c5cfa9d
performance tuning: prefer static simpset within functional closure of morphism (with notable impact on specifications "in" class, e.g. AFP/No_FTL_observers);
wenzelm
parents:
78065
diff
changeset
|
161 |
Morphism.set_context thy |
3fdf3c5cfa9d
performance tuning: prefer static simpset within functional closure of morphism (with notable impact on specifications "in" class, e.g. AFP/No_FTL_observers);
wenzelm
parents:
78065
diff
changeset
|
162 |
(base_morphism thy class $> |
3fdf3c5cfa9d
performance tuning: prefer static simpset within functional closure of morphism (with notable impact on specifications "in" class, e.g. AFP/No_FTL_observers);
wenzelm
parents:
78065
diff
changeset
|
163 |
Morphism.default (Element.eq_morphism ctxt (these_defs thy [class]))) |
3fdf3c5cfa9d
performance tuning: prefer static simpset within functional closure of morphism (with notable impact on specifications "in" class, e.g. AFP/No_FTL_observers);
wenzelm
parents:
78065
diff
changeset
|
164 |
end; |
47078 | 165 |
|
32850 | 166 |
val export_morphism = #export_morph oo the_class_data; |
28715
238f9966c80e
class morphism stemming from locale interpretation
haftmann
parents:
28674
diff
changeset
|
167 |
|
63347 | 168 |
fun pretty_param ctxt (c, ty) = |
169 |
Pretty.block |
|
170 |
[Name_Space.pretty ctxt (Proof_Context.const_space ctxt) c, Pretty.str " ::", |
|
171 |
Pretty.brk 1, Syntax.pretty_typ ctxt ty]; |
|
172 |
||
42359 | 173 |
fun print_classes ctxt = |
24218 | 174 |
let |
42360 | 175 |
val thy = Proof_Context.theory_of ctxt; |
24218 | 176 |
val algebra = Sign.classes_of thy; |
51510
b4f7e6734acc
tuned print_classes: more standard order, markup, formatting;
wenzelm
parents:
49687
diff
changeset
|
177 |
|
b4f7e6734acc
tuned print_classes: more standard order, markup, formatting;
wenzelm
parents:
49687
diff
changeset
|
178 |
val class_space = Proof_Context.class_space ctxt; |
b4f7e6734acc
tuned print_classes: more standard order, markup, formatting;
wenzelm
parents:
49687
diff
changeset
|
179 |
val type_space = Proof_Context.type_space ctxt; |
b4f7e6734acc
tuned print_classes: more standard order, markup, formatting;
wenzelm
parents:
49687
diff
changeset
|
180 |
|
24218 | 181 |
val arities = |
74232 | 182 |
Symtab.build (Sorts.arities_of algebra |> Symtab.fold (fn (tyco, arities) => |
183 |
fold (fn (class, _) => Symtab.map_default (class, []) (insert (op =) tyco)) arities)); |
|
51510
b4f7e6734acc
tuned print_classes: more standard order, markup, formatting;
wenzelm
parents:
49687
diff
changeset
|
184 |
|
b4f7e6734acc
tuned print_classes: more standard order, markup, formatting;
wenzelm
parents:
49687
diff
changeset
|
185 |
fun prt_supersort class = |
b4f7e6734acc
tuned print_classes: more standard order, markup, formatting;
wenzelm
parents:
49687
diff
changeset
|
186 |
Syntax.pretty_sort ctxt (Sign.minimize_sort thy (Sign.super_classes thy class)); |
b4f7e6734acc
tuned print_classes: more standard order, markup, formatting;
wenzelm
parents:
49687
diff
changeset
|
187 |
|
b4f7e6734acc
tuned print_classes: more standard order, markup, formatting;
wenzelm
parents:
49687
diff
changeset
|
188 |
fun prt_arity class tyco = |
24218 | 189 |
let |
190 |
val Ss = Sorts.mg_domain algebra tyco [class]; |
|
24920 | 191 |
in Syntax.pretty_arity ctxt (tyco, Ss, [class]) end; |
51510
b4f7e6734acc
tuned print_classes: more standard order, markup, formatting;
wenzelm
parents:
49687
diff
changeset
|
192 |
|
63347 | 193 |
fun prt_param (c, ty) = pretty_param ctxt (c, Type.strip_sorts_dummy ty); |
51510
b4f7e6734acc
tuned print_classes: more standard order, markup, formatting;
wenzelm
parents:
49687
diff
changeset
|
194 |
|
b4f7e6734acc
tuned print_classes: more standard order, markup, formatting;
wenzelm
parents:
49687
diff
changeset
|
195 |
fun prt_entry class = |
b4f7e6734acc
tuned print_classes: more standard order, markup, formatting;
wenzelm
parents:
49687
diff
changeset
|
196 |
Pretty.block |
55763 | 197 |
([Pretty.keyword1 "class", Pretty.brk 1, |
53539 | 198 |
Name_Space.pretty ctxt class_space class, Pretty.str ":", Pretty.fbrk, |
199 |
Pretty.block [Pretty.str "supersort: ", prt_supersort class]] @ |
|
68098 | 200 |
(case (these o Option.map #params o try (Axclass.get_info thy)) class of |
201 |
[] => [] |
|
202 |
| params => |
|
51510
b4f7e6734acc
tuned print_classes: more standard order, markup, formatting;
wenzelm
parents:
49687
diff
changeset
|
203 |
[Pretty.fbrk, Pretty.big_list "parameters:" (map prt_param params)]) @ |
68098 | 204 |
(case (these o Symtab.lookup arities) class of |
205 |
[] => [] |
|
206 |
| ars => |
|
51510
b4f7e6734acc
tuned print_classes: more standard order, markup, formatting;
wenzelm
parents:
49687
diff
changeset
|
207 |
[Pretty.fbrk, Pretty.big_list "instances:" |
b4f7e6734acc
tuned print_classes: more standard order, markup, formatting;
wenzelm
parents:
49687
diff
changeset
|
208 |
(map (prt_arity class) (sort (Name_Space.extern_ord ctxt type_space) ars))])); |
24218 | 209 |
in |
51510
b4f7e6734acc
tuned print_classes: more standard order, markup, formatting;
wenzelm
parents:
49687
diff
changeset
|
210 |
Sorts.all_classes algebra |
b4f7e6734acc
tuned print_classes: more standard order, markup, formatting;
wenzelm
parents:
49687
diff
changeset
|
211 |
|> sort (Name_Space.extern_ord ctxt class_space) |
b4f7e6734acc
tuned print_classes: more standard order, markup, formatting;
wenzelm
parents:
49687
diff
changeset
|
212 |
|> map prt_entry |
56334
6b3739fee456
some shortcuts for chunks, which sometimes avoid bulky string output;
wenzelm
parents:
56204
diff
changeset
|
213 |
|> Pretty.writeln_chunks2 |
24218 | 214 |
end; |
215 |
||
216 |
||
217 |
(* updaters *) |
|
218 |
||
32850 | 219 |
fun register class sups params base_sort base_morph export_morph |
52636 | 220 |
some_axiom some_assm_intro of_class thy = |
25002 | 221 |
let |
25368
f12613fda79d
proper implementation of check phase; non-qualified names for class operations
haftmann
parents:
25344
diff
changeset
|
222 |
val operations = map (fn (v_ty as (_, ty), (c, _)) => |
60346 | 223 |
(c, (class, ((ty, Free v_ty), false)))) params; |
25038 | 224 |
val add_class = Graph.new_node (class, |
59058
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents:
58837
diff
changeset
|
225 |
make_class_data (((map o apply2) fst params, base_sort, |
78064
4e865c45458b
clarified transfer / trim_context on persistent Token.source (e.g. attribute expressions): actually set/reset implicit context;
wenzelm
parents:
78060
diff
changeset
|
226 |
Morphism.reset_context base_morph, |
4e865c45458b
clarified transfer / trim_context on persistent Token.source (e.g. attribute expressions): actually set/reset implicit context;
wenzelm
parents:
78060
diff
changeset
|
227 |
Morphism.reset_context export_morph, |
4e865c45458b
clarified transfer / trim_context on persistent Token.source (e.g. attribute expressions): actually set/reset implicit context;
wenzelm
parents:
78060
diff
changeset
|
228 |
Option.map Thm.trim_context some_assm_intro, |
4e865c45458b
clarified transfer / trim_context on persistent Token.source (e.g. attribute expressions): actually set/reset implicit context;
wenzelm
parents:
78060
diff
changeset
|
229 |
Thm.trim_context of_class, |
4e865c45458b
clarified transfer / trim_context on persistent Token.source (e.g. attribute expressions): actually set/reset implicit context;
wenzelm
parents:
78060
diff
changeset
|
230 |
Option.map Thm.trim_context some_axiom), |
74152 | 231 |
(Thm.item_net, operations))) |
29526 | 232 |
#> fold (curry Graph.add_edge class) sups; |
46919 | 233 |
in Class_Data.map add_class thy end; |
24218 | 234 |
|
78453
3fdf3c5cfa9d
performance tuning: prefer static simpset within functional closure of morphism (with notable impact on specifications "in" class, e.g. AFP/No_FTL_observers);
wenzelm
parents:
78065
diff
changeset
|
235 |
fun activate_defs context class thms thy = |
3fdf3c5cfa9d
performance tuning: prefer static simpset within functional closure of morphism (with notable impact on specifications "in" class, e.g. AFP/No_FTL_observers);
wenzelm
parents:
78065
diff
changeset
|
236 |
(case Element.eq_morphism (Context.proof_of context) thms of |
68851
6c9825c1e26b
clarified signature: explicit type Locale.registration;
wenzelm
parents:
68351
diff
changeset
|
237 |
SOME eq_morph => |
78059 | 238 |
fold (fn cls => fn thy' => |
239 |
(Context.theory_map o Locale.amend_registration) |
|
240 |
{inst = (cls, base_morphism thy' cls), |
|
241 |
mixin = SOME (eq_morph, true), |
|
242 |
export = export_morphism thy' cls} thy') (heritage thy [class]) thy |
|
46919 | 243 |
| NONE => thy); |
29526 | 244 |
|
60346 | 245 |
fun register_operation class (c, t) input_only thy = |
25062 | 246 |
let |
29358 | 247 |
val base_sort = base_sort thy class; |
29439 | 248 |
val prep_typ = map_type_tfree |
249 |
(fn (v, sort) => if Name.aT = v |
|
250 |
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
|
251 |
val t' = map_types prep_typ t; |
f12613fda79d
proper implementation of check phase; non-qualified names for class operations
haftmann
parents:
25344
diff
changeset
|
252 |
val ty' = Term.fastype_of t'; |
25062 | 253 |
in |
254 |
thy |
|
57173
897cc57a6f55
always refine interpretation morphism using canonical constant's definition theorem
haftmann
parents:
57161
diff
changeset
|
255 |
|> (Class_Data.map o Graph.map_node class o map_class_data o apsnd o apsnd) |
60346 | 256 |
(cons (c, (class, ((ty', t'), input_only)))) |
57173
897cc57a6f55
always refine interpretation morphism using canonical constant's definition theorem
haftmann
parents:
57161
diff
changeset
|
257 |
end; |
897cc57a6f55
always refine interpretation morphism using canonical constant's definition theorem
haftmann
parents:
57161
diff
changeset
|
258 |
|
897cc57a6f55
always refine interpretation morphism using canonical constant's definition theorem
haftmann
parents:
57161
diff
changeset
|
259 |
fun register_def class def_thm thy = |
897cc57a6f55
always refine interpretation morphism using canonical constant's definition theorem
haftmann
parents:
57161
diff
changeset
|
260 |
let |
61085 | 261 |
val sym_thm = Thm.trim_context (Thm.symmetric def_thm) |
57173
897cc57a6f55
always refine interpretation morphism using canonical constant's definition theorem
haftmann
parents:
57161
diff
changeset
|
262 |
in |
897cc57a6f55
always refine interpretation morphism using canonical constant's definition theorem
haftmann
parents:
57161
diff
changeset
|
263 |
thy |
897cc57a6f55
always refine interpretation morphism using canonical constant's definition theorem
haftmann
parents:
57161
diff
changeset
|
264 |
|> (Class_Data.map o Graph.map_node class o map_class_data o apsnd o apfst) |
74150 | 265 |
(Item_Net.update sym_thm) |
78453
3fdf3c5cfa9d
performance tuning: prefer static simpset within functional closure of morphism (with notable impact on specifications "in" class, e.g. AFP/No_FTL_observers);
wenzelm
parents:
78065
diff
changeset
|
266 |
|> activate_defs (Context.Theory thy) class [sym_thm] |
25062 | 267 |
end; |
24218 | 268 |
|
269 |
||
24589 | 270 |
(** classes and class target **) |
24218 | 271 |
|
25002 | 272 |
(* class context syntax *) |
24748 | 273 |
|
60347 | 274 |
fun make_rewrite t c_ty = |
275 |
let |
|
73793
26c0ccf17f31
more accurate export morphism enables proper instantiation by interpretation
haftmann
parents:
72516
diff
changeset
|
276 |
val vs = strip_abs_vars t; |
60347 | 277 |
val vts = map snd vs |
278 |
|> Name.invent_names Name.context Name.uu |
|
279 |
|> map (fn (v, T) => Var ((v, 0), T)); |
|
280 |
in (betapplys (t, vts), betapplys (Const c_ty, vts)) end; |
|
281 |
||
60346 | 282 |
fun these_unchecks thy = |
60347 | 283 |
these_operations thy |
284 |
#> map_filter (fn (c, (_, ((ty, t), input_only))) => |
|
285 |
if input_only then NONE else SOME (make_rewrite t (c, ty))); |
|
286 |
||
287 |
fun these_unchecks_reversed thy = |
|
288 |
these_operations thy |
|
289 |
#> map (fn (c, (_, ((ty, t), _))) => (Const (c, ty), t)); |
|
29577 | 290 |
|
29632
c3d576157244
fixed reading of class specs: declare class operations in context
haftmann
parents:
29610
diff
changeset
|
291 |
fun redeclare_const thy c = |
30364
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents:
30344
diff
changeset
|
292 |
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
|
293 |
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
|
294 |
|
29577 | 295 |
fun synchronize_class_syntax sort base_sort ctxt = |
24914
95cda5dd58d5
added proper subclass concept; improved class target
haftmann
parents:
24901
diff
changeset
|
296 |
let |
42360 | 297 |
val thy = Proof_Context.theory_of ctxt; |
26596 | 298 |
val algebra = Sign.classes_of thy; |
29577 | 299 |
val operations = these_operations thy sort; |
26518 | 300 |
fun subst_class_typ sort = map_type_tfree (K (TVar ((Name.aT, 0), sort))); |
301 |
val primary_constraints = |
|
60346 | 302 |
(map o apsnd) (subst_class_typ base_sort o fst o fst o snd) operations; |
26518 | 303 |
val secondary_constraints = |
60346 | 304 |
(map o apsnd) (fn (class, ((ty, _), _)) => subst_class_typ [class] ty) operations; |
46919 | 305 |
fun improve (c, ty) = |
306 |
(case AList.lookup (op =) primary_constraints c of |
|
307 |
SOME ty' => |
|
308 |
(case try (Type.raw_match (ty', ty)) Vartab.empty of |
|
309 |
SOME tyenv => |
|
310 |
(case Vartab.lookup tyenv (Name.aT, 0) of |
|
311 |
SOME (_, ty' as TVar (vi, sort)) => |
|
312 |
if Type_Infer.is_param vi andalso Sorts.sort_le algebra (base_sort, sort) |
|
70387 | 313 |
then SOME (ty', Term.aT base_sort) |
46919 | 314 |
else NONE |
26238 | 315 |
| _ => NONE) |
316 |
| NONE => NONE) |
|
46919 | 317 |
| NONE => NONE); |
60346 | 318 |
fun subst (c, _) = Option.map (fst o snd) (AList.lookup (op =) operations c); |
29577 | 319 |
val unchecks = these_unchecks thy sort; |
25083 | 320 |
in |
321 |
ctxt |
|
29632
c3d576157244
fixed reading of class specs: declare class operations in context
haftmann
parents:
29610
diff
changeset
|
322 |
|> fold (redeclare_const thy o fst) primary_constraints |
60338 | 323 |
|> Overloading.map_improvable_syntax (K {primary_constraints = primary_constraints, |
324 |
secondary_constraints = secondary_constraints, improve = improve, subst = subst, |
|
325 |
no_subst_in_abbrev_mode = true, unchecks = unchecks}) |
|
26518 | 326 |
|> Overloading.set_primary_constraints |
25083 | 327 |
end; |
328 |
||
54866
7b9a67cbd48f
self-contained formulation of subclass command, avoiding hard-wired Named_Target.init
haftmann
parents:
54742
diff
changeset
|
329 |
fun synchronize_class_syntax_target class lthy = |
7b9a67cbd48f
self-contained formulation of subclass command, avoiding hard-wired Named_Target.init
haftmann
parents:
54742
diff
changeset
|
330 |
lthy |
7b9a67cbd48f
self-contained formulation of subclass command, avoiding hard-wired Named_Target.init
haftmann
parents:
54742
diff
changeset
|
331 |
|> Local_Theory.map_contexts |
7b9a67cbd48f
self-contained formulation of subclass command, avoiding hard-wired Named_Target.init
haftmann
parents:
54742
diff
changeset
|
332 |
(K (synchronize_class_syntax [class] (base_sort (Proof_Context.theory_of lthy) class))); |
7b9a67cbd48f
self-contained formulation of subclass command, avoiding hard-wired Named_Target.init
haftmann
parents:
54742
diff
changeset
|
333 |
|
29632
c3d576157244
fixed reading of class specs: declare class operations in context
haftmann
parents:
29610
diff
changeset
|
334 |
fun redeclare_operations thy sort = |
c3d576157244
fixed reading of class specs: declare class operations in context
haftmann
parents:
29610
diff
changeset
|
335 |
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
|
336 |
|
29577 | 337 |
fun begin sort base_sort ctxt = |
25083 | 338 |
ctxt |
70387 | 339 |
|> Variable.declare_term (Logic.mk_type (Term.aT base_sort)) |
29577 | 340 |
|> synchronize_class_syntax sort base_sort |
39378
df86b1b4ce10
more precise name for activation of improveable syntax
haftmann
parents:
39134
diff
changeset
|
341 |
|> Overloading.activate_improvable_syntax; |
24901
d3cbf79769b9
added first version of user-space type system for class target
haftmann
parents:
24847
diff
changeset
|
342 |
|
25311 | 343 |
fun init class thy = |
344 |
thy |
|
29509
1ff0f3f08a7b
migrated class package to new locale implementation
haftmann
parents:
29439
diff
changeset
|
345 |
|> Locale.init class |
29358 | 346 |
|> begin [class] (base_sort thy class); |
24914
95cda5dd58d5
added proper subclass concept; improved class target
haftmann
parents:
24901
diff
changeset
|
347 |
|
24748 | 348 |
|
27690
24738db98d34
some steps towards explicit class target for canonical interpretation
haftmann
parents:
27684
diff
changeset
|
349 |
(* class target *) |
24738db98d34
some steps towards explicit class target for canonical interpretation
haftmann
parents:
27684
diff
changeset
|
350 |
|
30364
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents:
30344
diff
changeset
|
351 |
val class_prefix = Logic.const_of_class o Long_Name.base_name; |
29526 | 352 |
|
57187
de00494fa8b3
less ad-hoc verision of educated guess: guess identity of declaration morphism wrt. canonical morphism rather than observing particular effects of declaration morphisms only;
haftmann
parents:
57186
diff
changeset
|
353 |
fun guess_morphism_identity (b, rhs) phi1 phi2 = |
de00494fa8b3
less ad-hoc verision of educated guess: guess identity of declaration morphism wrt. canonical morphism rather than observing particular effects of declaration morphisms only;
haftmann
parents:
57186
diff
changeset
|
354 |
let |
de00494fa8b3
less ad-hoc verision of educated guess: guess identity of declaration morphism wrt. canonical morphism rather than observing particular effects of declaration morphisms only;
haftmann
parents:
57186
diff
changeset
|
355 |
(*FIXME proper concept to identify morphism instead of educated guess*) |
58668 | 356 |
val name_of_binding = Name_Space.full_name Name_Space.global_naming; |
57187
de00494fa8b3
less ad-hoc verision of educated guess: guess identity of declaration morphism wrt. canonical morphism rather than observing particular effects of declaration morphisms only;
haftmann
parents:
57186
diff
changeset
|
357 |
val n1 = (name_of_binding o Morphism.binding phi1) b; |
de00494fa8b3
less ad-hoc verision of educated guess: guess identity of declaration morphism wrt. canonical morphism rather than observing particular effects of declaration morphisms only;
haftmann
parents:
57186
diff
changeset
|
358 |
val n2 = (name_of_binding o Morphism.binding phi2) b; |
de00494fa8b3
less ad-hoc verision of educated guess: guess identity of declaration morphism wrt. canonical morphism rather than observing particular effects of declaration morphisms only;
haftmann
parents:
57186
diff
changeset
|
359 |
val rhs1 = Morphism.term phi1 rhs; |
de00494fa8b3
less ad-hoc verision of educated guess: guess identity of declaration morphism wrt. canonical morphism rather than observing particular effects of declaration morphisms only;
haftmann
parents:
57186
diff
changeset
|
360 |
val rhs2 = Morphism.term phi2 rhs; |
de00494fa8b3
less ad-hoc verision of educated guess: guess identity of declaration morphism wrt. canonical morphism rather than observing particular effects of declaration morphisms only;
haftmann
parents:
57186
diff
changeset
|
361 |
in n1 = n2 andalso Term.aconv_untyped (rhs1, rhs2) end; |
de00494fa8b3
less ad-hoc verision of educated guess: guess identity of declaration morphism wrt. canonical morphism rather than observing particular effects of declaration morphisms only;
haftmann
parents:
57186
diff
changeset
|
362 |
|
60345
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
haftmann
parents:
60344
diff
changeset
|
363 |
fun target_const class phi0 prmode (b, rhs) lthy = |
57189
5140ddfccea7
re-unified approach towards class and locale consts, with refined terminology: foo_const_declaration denotes declaration for a particular logical layer, foo_const the full stack for a particular target
haftmann
parents:
57188
diff
changeset
|
364 |
let |
78065
11d6a1e62841
more careful treatment of set_context / reset_context for persistent morphisms;
wenzelm
parents:
78064
diff
changeset
|
365 |
val export = |
11d6a1e62841
more careful treatment of set_context / reset_context for persistent morphisms;
wenzelm
parents:
78064
diff
changeset
|
366 |
Morphism.set_context' lthy (Variable.export_morphism lthy (Local_Theory.target_of lthy)); |
60345
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
haftmann
parents:
60344
diff
changeset
|
367 |
val guess_identity = guess_morphism_identity (b, rhs) export; |
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
haftmann
parents:
60344
diff
changeset
|
368 |
val guess_canonical = guess_morphism_identity (b, rhs) (export $> phi0); |
57189
5140ddfccea7
re-unified approach towards class and locale consts, with refined terminology: foo_const_declaration denotes declaration for a particular logical layer, foo_const the full stack for a particular target
haftmann
parents:
57188
diff
changeset
|
369 |
in |
60345
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
haftmann
parents:
60344
diff
changeset
|
370 |
lthy |
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
haftmann
parents:
60344
diff
changeset
|
371 |
|> Generic_Target.locale_target_const class |
57189
5140ddfccea7
re-unified approach towards class and locale consts, with refined terminology: foo_const_declaration denotes declaration for a particular logical layer, foo_const the full stack for a particular target
haftmann
parents:
57188
diff
changeset
|
372 |
(not o (guess_identity orf guess_canonical)) prmode ((b, NoSyn), rhs) |
5140ddfccea7
re-unified approach towards class and locale consts, with refined terminology: foo_const_declaration denotes declaration for a particular logical layer, foo_const the full stack for a particular target
haftmann
parents:
57188
diff
changeset
|
373 |
end; |
57072
dfac6ef0ca28
moved const declaration further down in bootstrap hierarchy: keep Named_Target free of low-level stuff
haftmann
parents:
57070
diff
changeset
|
374 |
|
60345
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
haftmann
parents:
60344
diff
changeset
|
375 |
local |
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
haftmann
parents:
60344
diff
changeset
|
376 |
|
57119 | 377 |
fun dangling_params_for lthy class (type_params, term_params) = |
27690
24738db98d34
some steps towards explicit class target for canonical interpretation
haftmann
parents:
27684
diff
changeset
|
378 |
let |
57119 | 379 |
val class_param_names = |
380 |
map fst (these_params (Proof_Context.theory_of lthy) [class]); |
|
381 |
val dangling_term_params = |
|
382 |
subtract (fn (v, Free (w, _)) => v = w | _ => false) class_param_names term_params; |
|
57161
6254c51cd210
formal treatment of dangling parameters for class abbrevs analogously to class consts
haftmann
parents:
57148
diff
changeset
|
383 |
in (type_params, dangling_term_params) end; |
57119 | 384 |
|
385 |
fun global_def (b, eq) thy = |
|
68350 | 386 |
let |
387 |
val ((_, def_thm), thy') = thy |> Thm.add_def_global false false (b, eq); |
|
68351
bcdc4c21ab1d
varify frees, notably dangling_params (see also e0cd57aeb60c);
wenzelm
parents:
68350
diff
changeset
|
388 |
val def_thm' = def_thm |> Thm.forall_intr_frees |> Thm.forall_elim_vars 0 |> Thm.varifyT_global; |
68350 | 389 |
val (_, thy'') = thy' |> Global_Theory.store_thm (b, def_thm'); |
390 |
in (def_thm', thy'') end; |
|
57119 | 391 |
|
57142 | 392 |
fun canonical_const class phi dangling_params ((b, mx), rhs) thy = |
57119 | 393 |
let |
57147 | 394 |
val b_def = Binding.suffix_name "_dict" b; |
395 |
val c = Sign.full_name thy b; |
|
396 |
val ty = map Term.fastype_of dangling_params ---> Term.fastype_of rhs; |
|
397 |
val def_eq = Logic.mk_equals (list_comb (Const (c, ty), dangling_params), rhs) |
|
29577 | 398 |
|> map_types Type.strip_sorts; |
27690
24738db98d34
some steps towards explicit class target for canonical interpretation
haftmann
parents:
27684
diff
changeset
|
399 |
in |
29577 | 400 |
thy |
57147 | 401 |
|> Sign.declare_const_global ((b, Type.strip_sorts ty), mx) |
29577 | 402 |
|> snd |
57119 | 403 |
|> global_def (b_def, def_eq) |
57173
897cc57a6f55
always refine interpretation morphism using canonical constant's definition theorem
haftmann
parents:
57161
diff
changeset
|
404 |
|-> (fn def_thm => register_def class def_thm) |
60346 | 405 |
|> null dangling_params ? register_operation class (c, rhs) false |
57147 | 406 |
|> Sign.add_const_constraint (c, SOME ty) |
27690
24738db98d34
some steps towards explicit class target for canonical interpretation
haftmann
parents:
27684
diff
changeset
|
407 |
end; |
24738db98d34
some steps towards explicit class target for canonical interpretation
haftmann
parents:
27684
diff
changeset
|
408 |
|
38619
25e401d53900
tuned: less formal noise in Named_Target, more coherence in Class
haftmann
parents:
38392
diff
changeset
|
409 |
in |
25e401d53900
tuned: less formal noise in Named_Target, more coherence in Class
haftmann
parents:
38392
diff
changeset
|
410 |
|
57119 | 411 |
fun const class ((b, mx), lhs) params lthy = |
412 |
let |
|
78453
3fdf3c5cfa9d
performance tuning: prefer static simpset within functional closure of morphism (with notable impact on specifications "in" class, e.g. AFP/No_FTL_observers);
wenzelm
parents:
78065
diff
changeset
|
413 |
val phi = morphism lthy class; |
57161
6254c51cd210
formal treatment of dangling parameters for class abbrevs analogously to class consts
haftmann
parents:
57148
diff
changeset
|
414 |
val dangling_params = map (Morphism.term phi) (uncurry append (dangling_params_for lthy class params)); |
57119 | 415 |
in |
416 |
lthy |
|
60337 | 417 |
|> target_const class phi Syntax.mode_default (b, lhs) |
57148
4069c9b3803a
definition in class: provide explicit auxiliary abbreviation carrying potential mixfix syntax in presence of dangling parameters
haftmann
parents:
57147
diff
changeset
|
418 |
|> Local_Theory.raw_theory (canonical_const class phi dangling_params |
4069c9b3803a
definition in class: provide explicit auxiliary abbreviation carrying potential mixfix syntax in presence of dangling parameters
haftmann
parents:
57147
diff
changeset
|
419 |
((Morphism.binding phi b, if null dangling_params then mx else NoSyn), Morphism.term phi lhs)) |
57192
180e955711cf
yet another attempt for terminology: foo_target_bar denotes an operation bar operating solely on the target context of target foo, foo_bar denotes a whole stack of operations to accomplish bar for target foo
haftmann
parents:
57189
diff
changeset
|
420 |
|> Generic_Target.standard_const (fn (this, other) => other <> 0 andalso this <> other) |
57148
4069c9b3803a
definition in class: provide explicit auxiliary abbreviation carrying potential mixfix syntax in presence of dangling parameters
haftmann
parents:
57147
diff
changeset
|
421 |
Syntax.mode_default ((b, if null dangling_params then NoSyn else mx), lhs) |
57141 | 422 |
|> synchronize_class_syntax_target class |
57119 | 423 |
end; |
57072
dfac6ef0ca28
moved const declaration further down in bootstrap hierarchy: keep Named_Target free of low-level stuff
haftmann
parents:
57070
diff
changeset
|
424 |
|
60345
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
haftmann
parents:
60344
diff
changeset
|
425 |
end; |
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
haftmann
parents:
60344
diff
changeset
|
426 |
|
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
haftmann
parents:
60344
diff
changeset
|
427 |
local |
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
haftmann
parents:
60344
diff
changeset
|
428 |
|
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
haftmann
parents:
60344
diff
changeset
|
429 |
fun canonical_abbrev class phi prmode with_syntax ((b, mx), rhs) thy = |
57141 | 430 |
let |
60345
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
haftmann
parents:
60344
diff
changeset
|
431 |
val c = Sign.full_name thy b; |
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
haftmann
parents:
60344
diff
changeset
|
432 |
val constrain = map_atyps (fn T as TFree (v, _) => |
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
haftmann
parents:
60344
diff
changeset
|
433 |
if v = Name.aT then TFree (v, [class]) else T | T => T); |
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
haftmann
parents:
60344
diff
changeset
|
434 |
val rhs' = map_types constrain rhs; |
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
haftmann
parents:
60344
diff
changeset
|
435 |
in |
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
haftmann
parents:
60344
diff
changeset
|
436 |
thy |
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
haftmann
parents:
60344
diff
changeset
|
437 |
|> Sign.add_abbrev (#1 prmode) (b, Logic.varify_types_global rhs') |
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
haftmann
parents:
60344
diff
changeset
|
438 |
|> snd |
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
haftmann
parents:
60344
diff
changeset
|
439 |
|> with_syntax ? Sign.notation true prmode [(Const (c, fastype_of rhs), mx)] |
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
haftmann
parents:
60344
diff
changeset
|
440 |
|> with_syntax ? register_operation class (c, rhs) |
60346 | 441 |
(#1 prmode = Print_Mode.input) |
60345
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
haftmann
parents:
60344
diff
changeset
|
442 |
|> Sign.add_const_constraint (c, SOME (fastype_of rhs')) |
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
haftmann
parents:
60344
diff
changeset
|
443 |
end; |
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
haftmann
parents:
60344
diff
changeset
|
444 |
|
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
haftmann
parents:
60344
diff
changeset
|
445 |
fun canonical_abbrev_target class phi prmode ((b, mx), rhs) lthy = |
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
haftmann
parents:
60344
diff
changeset
|
446 |
let |
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
haftmann
parents:
60344
diff
changeset
|
447 |
val thy = Proof_Context.theory_of lthy; |
60347 | 448 |
val preprocess = perhaps (try (Pattern.rewrite_term_top thy (these_unchecks thy [class]) [])); |
73793
26c0ccf17f31
more accurate export morphism enables proper instantiation by interpretation
haftmann
parents:
72516
diff
changeset
|
449 |
val (global_rhs, (_, (_, term_params))) = |
60345
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
haftmann
parents:
60344
diff
changeset
|
450 |
Generic_Target.export_abbrev lthy preprocess rhs; |
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
haftmann
parents:
60344
diff
changeset
|
451 |
val mx' = Generic_Target.check_mixfix_global (b, null term_params) mx; |
57141 | 452 |
in |
453 |
lthy |
|
60345
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
haftmann
parents:
60344
diff
changeset
|
454 |
|> Local_Theory.raw_theory (canonical_abbrev class phi prmode (null term_params) |
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
haftmann
parents:
60344
diff
changeset
|
455 |
((Morphism.binding phi b, mx'), Logic.unvarify_types_global global_rhs)) |
57141 | 456 |
end; |
38619
25e401d53900
tuned: less formal noise in Named_Target, more coherence in Class
haftmann
parents:
38392
diff
changeset
|
457 |
|
60345
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
haftmann
parents:
60344
diff
changeset
|
458 |
fun further_abbrev_target class phi prmode (b, mx) rhs params = |
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
haftmann
parents:
60344
diff
changeset
|
459 |
Generic_Target.background_abbrev (b, rhs) (snd params) |
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
haftmann
parents:
60344
diff
changeset
|
460 |
#-> (fn (lhs, _) => target_const class phi prmode (b, lhs) |
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
haftmann
parents:
60344
diff
changeset
|
461 |
#> Generic_Target.standard_const (fn (this, other) => other <> 0 andalso this <> other) prmode ((b, mx), lhs)) |
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
haftmann
parents:
60344
diff
changeset
|
462 |
|
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
haftmann
parents:
60344
diff
changeset
|
463 |
in |
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
haftmann
parents:
60344
diff
changeset
|
464 |
|
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
haftmann
parents:
60344
diff
changeset
|
465 |
fun abbrev class prmode ((b, mx), rhs) lthy = |
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
haftmann
parents:
60344
diff
changeset
|
466 |
let |
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
haftmann
parents:
60344
diff
changeset
|
467 |
val thy = Proof_Context.theory_of lthy; |
78453
3fdf3c5cfa9d
performance tuning: prefer static simpset within functional closure of morphism (with notable impact on specifications "in" class, e.g. AFP/No_FTL_observers);
wenzelm
parents:
78065
diff
changeset
|
468 |
val phi = morphism lthy class; |
60345
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
haftmann
parents:
60344
diff
changeset
|
469 |
val rhs_generic = |
60347 | 470 |
perhaps (try (Pattern.rewrite_term_top thy (these_unchecks_reversed thy [class]) [])) rhs; |
60345
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
haftmann
parents:
60344
diff
changeset
|
471 |
in |
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
haftmann
parents:
60344
diff
changeset
|
472 |
lthy |
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
haftmann
parents:
60344
diff
changeset
|
473 |
|> canonical_abbrev_target class phi prmode ((b, mx), rhs) |
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
haftmann
parents:
60344
diff
changeset
|
474 |
|> Generic_Target.abbrev (further_abbrev_target class phi) prmode ((b, mx), rhs_generic) |
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
haftmann
parents:
60344
diff
changeset
|
475 |
||> synchronize_class_syntax_target class |
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
haftmann
parents:
60344
diff
changeset
|
476 |
end; |
60344
a40369ea3ba5
self-contained formulation of abbrev for named targets
haftmann
parents:
60341
diff
changeset
|
477 |
|
38619
25e401d53900
tuned: less formal noise in Named_Target, more coherence in Class
haftmann
parents:
38392
diff
changeset
|
478 |
end; |
25e401d53900
tuned: less formal noise in Named_Target, more coherence in Class
haftmann
parents:
38392
diff
changeset
|
479 |
|
27690
24738db98d34
some steps towards explicit class target for canonical interpretation
haftmann
parents:
27684
diff
changeset
|
480 |
|
54866
7b9a67cbd48f
self-contained formulation of subclass command, avoiding hard-wired Named_Target.init
haftmann
parents:
54742
diff
changeset
|
481 |
(* subclasses *) |
7b9a67cbd48f
self-contained formulation of subclass command, avoiding hard-wired Named_Target.init
haftmann
parents:
54742
diff
changeset
|
482 |
|
7b9a67cbd48f
self-contained formulation of subclass command, avoiding hard-wired Named_Target.init
haftmann
parents:
54742
diff
changeset
|
483 |
fun register_subclass (sub, sup) some_dep_morph some_witn export lthy = |
7b9a67cbd48f
self-contained formulation of subclass command, avoiding hard-wired Named_Target.init
haftmann
parents:
54742
diff
changeset
|
484 |
let |
7b9a67cbd48f
self-contained formulation of subclass command, avoiding hard-wired Named_Target.init
haftmann
parents:
54742
diff
changeset
|
485 |
val thy = Proof_Context.theory_of lthy; |
78041 | 486 |
val conclude_witness = |
487 |
Thm.trim_context o Drule.export_without_context_open o Element.conclude_witness lthy; |
|
488 |
val intros = |
|
489 |
(snd o rules thy) sup :: |
|
490 |
map_filter I [Option.map conclude_witness some_witn, (fst o rules thy) sub]; |
|
54866
7b9a67cbd48f
self-contained formulation of subclass command, avoiding hard-wired Named_Target.init
haftmann
parents:
54742
diff
changeset
|
491 |
val classrel = |
7b9a67cbd48f
self-contained formulation of subclass command, avoiding hard-wired Named_Target.init
haftmann
parents:
54742
diff
changeset
|
492 |
Goal.prove_sorry_global thy [] [] (Logic.mk_classrel (sub, sup)) |
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
59423
diff
changeset
|
493 |
(fn {context = ctxt, ...} => EVERY (map (TRYALL o resolve_tac ctxt o single) intros)); |
54866
7b9a67cbd48f
self-contained formulation of subclass command, avoiding hard-wired Named_Target.init
haftmann
parents:
54742
diff
changeset
|
494 |
val diff_sort = Sign.complete_sort thy [sup] |
7b9a67cbd48f
self-contained formulation of subclass command, avoiding hard-wired Named_Target.init
haftmann
parents:
54742
diff
changeset
|
495 |
|> subtract (op =) (Sign.complete_sort thy [sub]) |
7b9a67cbd48f
self-contained formulation of subclass command, avoiding hard-wired Named_Target.init
haftmann
parents:
54742
diff
changeset
|
496 |
|> filter (is_class thy); |
69048
f79aeac59e15
tuned -- removed spurious dead code from 7b9a67cbd48f;
wenzelm
parents:
69017
diff
changeset
|
497 |
val add_dependency = |
68851
6c9825c1e26b
clarified signature: explicit type Locale.registration;
wenzelm
parents:
68351
diff
changeset
|
498 |
(case some_dep_morph of |
6c9825c1e26b
clarified signature: explicit type Locale.registration;
wenzelm
parents:
68351
diff
changeset
|
499 |
SOME dep_morph => |
73845 | 500 |
Generic_Target.locale_dependency sub |
69058 | 501 |
{inst = (sup, dep_morph $> Element.satisfy_morphism (the_list some_witn)), |
68851
6c9825c1e26b
clarified signature: explicit type Locale.registration;
wenzelm
parents:
68351
diff
changeset
|
502 |
mixin = NONE, export = export} |
6c9825c1e26b
clarified signature: explicit type Locale.registration;
wenzelm
parents:
68351
diff
changeset
|
503 |
| NONE => I); |
54866
7b9a67cbd48f
self-contained formulation of subclass command, avoiding hard-wired Named_Target.init
haftmann
parents:
54742
diff
changeset
|
504 |
in |
7b9a67cbd48f
self-contained formulation of subclass command, avoiding hard-wired Named_Target.init
haftmann
parents:
54742
diff
changeset
|
505 |
lthy |
7b9a67cbd48f
self-contained formulation of subclass command, avoiding hard-wired Named_Target.init
haftmann
parents:
54742
diff
changeset
|
506 |
|> Local_Theory.raw_theory |
7b9a67cbd48f
self-contained formulation of subclass command, avoiding hard-wired Named_Target.init
haftmann
parents:
54742
diff
changeset
|
507 |
(Axclass.add_classrel classrel |
7b9a67cbd48f
self-contained formulation of subclass command, avoiding hard-wired Named_Target.init
haftmann
parents:
54742
diff
changeset
|
508 |
#> Class_Data.map (Graph.add_edge (sub, sup)) |
78453
3fdf3c5cfa9d
performance tuning: prefer static simpset within functional closure of morphism (with notable impact on specifications "in" class, e.g. AFP/No_FTL_observers);
wenzelm
parents:
78065
diff
changeset
|
509 |
#> activate_defs (Context.Proof lthy) sub (these_defs thy diff_sort)) |
69048
f79aeac59e15
tuned -- removed spurious dead code from 7b9a67cbd48f;
wenzelm
parents:
69017
diff
changeset
|
510 |
|> add_dependency |
54866
7b9a67cbd48f
self-contained formulation of subclass command, avoiding hard-wired Named_Target.init
haftmann
parents:
54742
diff
changeset
|
511 |
|> synchronize_class_syntax_target sub |
7b9a67cbd48f
self-contained formulation of subclass command, avoiding hard-wired Named_Target.init
haftmann
parents:
54742
diff
changeset
|
512 |
end; |
31635 | 513 |
|
514 |
local |
|
515 |
||
516 |
fun gen_classrel mk_prop classrel thy = |
|
517 |
let |
|
518 |
fun after_qed results = |
|
51685
385ef6706252
more standard module name Axclass (according to file name);
wenzelm
parents:
51578
diff
changeset
|
519 |
Proof_Context.background_theory ((fold o fold) Axclass.add_classrel results); |
31635 | 520 |
in |
521 |
thy |
|
42360 | 522 |
|> Proof_Context.init_global |
36323
655e2d74de3a
modernized naming conventions of main Isar proof elements;
wenzelm
parents:
36106
diff
changeset
|
523 |
|> Proof.theorem NONE after_qed [[(mk_prop thy classrel, [])]] |
31635 | 524 |
end; |
525 |
||
526 |
in |
|
527 |
||
528 |
val classrel = |
|
51685
385ef6706252
more standard module name Axclass (according to file name);
wenzelm
parents:
51578
diff
changeset
|
529 |
gen_classrel (Logic.mk_classrel oo Axclass.cert_classrel); |
31635 | 530 |
val classrel_cmd = |
51685
385ef6706252
more standard module name Axclass (according to file name);
wenzelm
parents:
51578
diff
changeset
|
531 |
gen_classrel (Logic.mk_classrel oo Axclass.read_classrel); |
31635 | 532 |
|
533 |
end; (*local*) |
|
534 |
||
535 |
||
25462 | 536 |
(** instantiation target **) |
537 |
||
538 |
(* bookkeeping *) |
|
539 |
||
540 |
datatype instantiation = Instantiation of { |
|
25864 | 541 |
arities: string list * (string * sort) list * sort, |
25462 | 542 |
params: ((string * string) * (string * typ)) list |
25603 | 543 |
(*(instantiation parameter, type constructor), (local instantiation parameter, typ)*) |
25462 | 544 |
} |
545 |
||
59150 | 546 |
fun make_instantiation (arities, params) = |
547 |
Instantiation {arities = arities, params = params}; |
|
548 |
||
549 |
val empty_instantiation = make_instantiation (([], [], []), []); |
|
550 |
||
33519 | 551 |
structure Instantiation = Proof_Data |
25462 | 552 |
( |
46919 | 553 |
type T = instantiation; |
59150 | 554 |
fun init _ = empty_instantiation; |
25462 | 555 |
); |
556 |
||
46919 | 557 |
val get_instantiation = |
558 |
(fn Instantiation data => data) o Instantiation.get o Local_Theory.target_of; |
|
25462 | 559 |
|
46919 | 560 |
fun map_instantiation f = |
561 |
(Local_Theory.target o Instantiation.map) |
|
59150 | 562 |
(fn Instantiation {arities, params} => make_instantiation (f (arities, params))); |
46919 | 563 |
|
564 |
fun the_instantiation lthy = |
|
565 |
(case get_instantiation lthy of |
|
566 |
{arities = ([], [], []), ...} => error "No instantiation target" |
|
567 |
| data => data); |
|
25462 | 568 |
|
25485 | 569 |
val instantiation_params = #params o get_instantiation; |
25462 | 570 |
|
30519
c05c0199826f
coherent binding policy with primitive target operations
haftmann
parents:
30364
diff
changeset
|
571 |
fun instantiation_param lthy b = instantiation_params lthy |
c05c0199826f
coherent binding policy with primitive target operations
haftmann
parents:
30364
diff
changeset
|
572 |
|> find_first (fn (_, (v, _)) => Binding.name_of b = v) |
25462 | 573 |
|> Option.map (fst o fst); |
574 |
||
31869 | 575 |
fun read_multi_arity thy (raw_tycos, raw_sorts, raw_sort) = |
576 |
let |
|
42360 | 577 |
val ctxt = Proof_Context.init_global thy; |
578 |
val all_arities = map (fn raw_tyco => Proof_Context.read_arity ctxt |
|
31869 | 579 |
(raw_tyco, raw_sorts, raw_sort)) raw_tycos; |
580 |
val tycos = map #1 all_arities; |
|
581 |
val (_, sorts, sort) = hd all_arities; |
|
43329
84472e198515
tuned signature: Name.invent and Name.invent_names;
wenzelm
parents:
42402
diff
changeset
|
582 |
val vs = Name.invent_names Name.context Name.aT sorts; |
31869 | 583 |
in (tycos, vs, sort) end; |
584 |
||
25462 | 585 |
|
586 |
(* syntax *) |
|
587 |
||
26238 | 588 |
fun synchronize_inst_syntax ctxt = |
25462 | 589 |
let |
46919 | 590 |
val Instantiation {params, ...} = Instantiation.get ctxt; |
31249 | 591 |
|
51685
385ef6706252
more standard module name Axclass (according to file name);
wenzelm
parents:
51578
diff
changeset
|
592 |
val lookup_inst_param = Axclass.lookup_inst_param |
42360 | 593 |
(Sign.consts_of (Proof_Context.theory_of ctxt)) params; |
46919 | 594 |
fun subst (c, ty) = |
595 |
(case lookup_inst_param (c, ty) of |
|
596 |
SOME (v_ty as (_, ty)) => SOME (ty, Free v_ty) |
|
597 |
| NONE => NONE); |
|
26238 | 598 |
val unchecks = |
599 |
map (fn ((c, _), v_ty as (_, ty)) => (Free v_ty, Const (c, ty))) params; |
|
600 |
in |
|
601 |
ctxt |
|
60338 | 602 |
|> Overloading.map_improvable_syntax (fn {primary_constraints, improve, ...} => |
603 |
{primary_constraints = primary_constraints, secondary_constraints = [], |
|
604 |
improve = improve, subst = subst, no_subst_in_abbrev_mode = false, |
|
605 |
unchecks = unchecks}) |
|
26238 | 606 |
end; |
25462 | 607 |
|
42385
b46b47775cbe
simplified Sorts.class_error: plain Proof.context;
wenzelm
parents:
42383
diff
changeset
|
608 |
fun resort_terms ctxt algebra consts constraints ts = |
38382 | 609 |
let |
42385
b46b47775cbe
simplified Sorts.class_error: plain Proof.context;
wenzelm
parents:
42383
diff
changeset
|
610 |
fun matchings (Const (c_ty as (c, _))) = |
b46b47775cbe
simplified Sorts.class_error: plain Proof.context;
wenzelm
parents:
42383
diff
changeset
|
611 |
(case constraints c of |
b46b47775cbe
simplified Sorts.class_error: plain Proof.context;
wenzelm
parents:
42383
diff
changeset
|
612 |
NONE => I |
b46b47775cbe
simplified Sorts.class_error: plain Proof.context;
wenzelm
parents:
42383
diff
changeset
|
613 |
| SOME sorts => |
b46b47775cbe
simplified Sorts.class_error: plain Proof.context;
wenzelm
parents:
42383
diff
changeset
|
614 |
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
|
615 |
| matchings _ = I; |
74232 | 616 |
val tvartab = Vartab.build ((fold o fold_aterms) matchings ts) |
61262
7bd1eb4b056e
tuned signature: eliminated pointless type Context.pretty;
wenzelm
parents:
61085
diff
changeset
|
617 |
handle Sorts.CLASS_ERROR e => error (Sorts.class_error (Context.Proof ctxt) e); |
38382 | 618 |
val inst = map_type_tvar |
619 |
(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
|
620 |
in if Vartab.is_empty tvartab then ts else (map o map_types) inst ts end; |
38382 | 621 |
|
25462 | 622 |
|
623 |
(* target *) |
|
624 |
||
63347 | 625 |
fun define_overloaded (c, U) b (b_def, rhs) lthy = |
626 |
let |
|
627 |
val name = Binding.name_of b; |
|
628 |
val pos = Binding.pos_of b; |
|
629 |
val _ = |
|
630 |
if Context_Position.is_reported lthy pos then |
|
631 |
Position.report_text pos Markup.class_parameter |
|
632 |
(Pretty.string_of |
|
633 |
(Pretty.block [Pretty.keyword1 "class", Pretty.brk 1, |
|
634 |
Pretty.str "parameter", Pretty.brk 1, pretty_param lthy (c, U)])) |
|
635 |
else (); |
|
636 |
in |
|
637 |
lthy |> Local_Theory.background_theory_result |
|
638 |
(Axclass.declare_overloaded (c, U) ##>> Axclass.define_overloaded b_def (c, rhs)) |
|
639 |
||> (map_instantiation o apsnd) (filter_out (fn (_, (v', _)) => v' = name)) |
|
640 |
||> Local_Theory.map_contexts (K synchronize_inst_syntax) |
|
641 |
end; |
|
38382 | 642 |
|
47289 | 643 |
fun foundation (((b, U), mx), (b_def, rhs)) params lthy = |
46919 | 644 |
(case instantiation_param lthy b of |
645 |
SOME c => |
|
63347 | 646 |
if Mixfix.is_empty mx then lthy |> define_overloaded (c, U) b (b_def, rhs) |
62765
5b95a12b7b19
tuned messages -- position is usually missing here;
wenzelm
parents:
62752
diff
changeset
|
647 |
else error ("Illegal mixfix syntax for overloaded constant " ^ quote c) |
60341
fcbd7f0c52c3
clearly separated target primitives (target_foo) from self-contained target operations (foo)
haftmann
parents:
60338
diff
changeset
|
648 |
| NONE => lthy |> Generic_Target.theory_target_foundation (((b, U), mx), (b_def, rhs)) params); |
38382 | 649 |
|
650 |
fun pretty lthy = |
|
26518 | 651 |
let |
46919 | 652 |
val {arities = (tycos, vs, sort), params} = the_instantiation lthy; |
38382 | 653 |
fun pr_arity tyco = Syntax.pretty_arity lthy (tyco, map snd vs, sort); |
654 |
fun pr_param ((c, _), (v, ty)) = |
|
42359 | 655 |
Pretty.block (Pretty.breaks |
55304 | 656 |
[Pretty.str v, Pretty.str "==", Proof_Context.pretty_const lthy c, |
42359 | 657 |
Pretty.str "::", Syntax.pretty_typ lthy ty]); |
60246 | 658 |
in |
659 |
[Pretty.block |
|
660 |
(Pretty.fbreaks (Pretty.keyword1 "instantiation" :: map pr_arity tycos @ map pr_param params))] |
|
661 |
end; |
|
26518 | 662 |
|
38382 | 663 |
fun conclude lthy = |
664 |
let |
|
42359 | 665 |
val (tycos, vs, sort) = #arities (the_instantiation lthy); |
42360 | 666 |
val thy = Proof_Context.theory_of lthy; |
42359 | 667 |
val _ = tycos |> List.app (fn tyco => |
46919 | 668 |
if Sign.of_sort thy (Type (tyco, map TFree vs), sort) then () |
55304 | 669 |
else error ("Missing instance proof for type " ^ quote (Proof_Context.markup_type lthy tyco))); |
38382 | 670 |
in lthy end; |
671 |
||
73793
26c0ccf17f31
more accurate export morphism enables proper instantiation by interpretation
haftmann
parents:
72516
diff
changeset
|
672 |
fun registration thy_ctxt {inst, mixin, export} lthy = |
26c0ccf17f31
more accurate export morphism enables proper instantiation by interpretation
haftmann
parents:
72516
diff
changeset
|
673 |
lthy |
73845 | 674 |
|> Generic_Target.theory_registration |
73793
26c0ccf17f31
more accurate export morphism enables proper instantiation by interpretation
haftmann
parents:
72516
diff
changeset
|
675 |
{inst = inst, |
26c0ccf17f31
more accurate export morphism enables proper instantiation by interpretation
haftmann
parents:
72516
diff
changeset
|
676 |
mixin = mixin, |
26c0ccf17f31
more accurate export morphism enables proper instantiation by interpretation
haftmann
parents:
72516
diff
changeset
|
677 |
export = export $> Proof_Context.export_morphism lthy thy_ctxt} |
26c0ccf17f31
more accurate export morphism enables proper instantiation by interpretation
haftmann
parents:
72516
diff
changeset
|
678 |
(*handle fixed types variables on target context properly*); |
26c0ccf17f31
more accurate export morphism enables proper instantiation by interpretation
haftmann
parents:
72516
diff
changeset
|
679 |
|
38382 | 680 |
fun instantiation (tycos, vs, sort) thy = |
25462 | 681 |
let |
25536 | 682 |
val _ = if null tycos then error "At least one arity must be given" else (); |
51685
385ef6706252
more standard module name Axclass (according to file name);
wenzelm
parents:
51578
diff
changeset
|
683 |
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
|
684 |
fun get_param tyco (param, (_, (c, ty))) = |
51685
385ef6706252
more standard module name Axclass (according to file name);
wenzelm
parents:
51578
diff
changeset
|
685 |
if can (Axclass.param_of_inst thy) (c, tyco) |
25603 | 686 |
then NONE else SOME ((c, tyco), |
46917 | 687 |
(param ^ "_" ^ Long_Name.base_name tyco, map_atyps (K (Type (tyco, map TFree vs))) ty)); |
31249 | 688 |
val params = map_product get_param tycos class_params |> map_filter I; |
51578 | 689 |
val _ = if null params andalso forall (fn tyco => can (Sign.arity_sorts thy tyco) sort) tycos |
690 |
then error "No parameters and no pending instance proof obligations in instantiation." |
|
691 |
else (); |
|
26518 | 692 |
val primary_constraints = map (apsnd |
31249 | 693 |
(map_atyps (K (TVar ((Name.aT, 0), [])))) o snd o snd) class_params; |
26518 | 694 |
val algebra = Sign.classes_of thy |
61262
7bd1eb4b056e
tuned signature: eliminated pointless type Context.pretty;
wenzelm
parents:
61085
diff
changeset
|
695 |
|> fold (fn tyco => Sorts.add_arities (Context.Theory thy) |
26518 | 696 |
(tyco, map (fn class => (class, map snd vs)) sort)) tycos; |
697 |
val consts = Sign.consts_of thy; |
|
698 |
val improve_constraints = AList.lookup (op =) |
|
31249 | 699 |
(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
|
700 |
fun resort_check ctxt ts = resort_terms ctxt algebra consts improve_constraints ts; |
51685
385ef6706252
more standard module name Axclass (according to file name);
wenzelm
parents:
51578
diff
changeset
|
701 |
val lookup_inst_param = Axclass.lookup_inst_param consts params; |
42388 | 702 |
fun improve (c, ty) = |
703 |
(case lookup_inst_param (c, ty) of |
|
704 |
SOME (_, ty') => if Sign.typ_instance thy (ty', ty) then SOME (ty, ty') else NONE |
|
705 |
| NONE => NONE); |
|
25485 | 706 |
in |
707 |
thy |
|
72516
17dc99589a91
unified Local_Theory.init with Generic_Target.init
haftmann
parents:
72505
diff
changeset
|
708 |
|> Local_Theory.init |
66335
a849ce33923d
treat exit separate from regular local theory operations
haftmann
parents:
63347
diff
changeset
|
709 |
{background_naming = Sign.naming_of thy, |
66337
5caea089dd17
more structural sharing between common target Generic_Target.init
haftmann
parents:
66335
diff
changeset
|
710 |
setup = Proof_Context.init_global |
5caea089dd17
more structural sharing between common target Generic_Target.init
haftmann
parents:
66335
diff
changeset
|
711 |
#> Instantiation.put (make_instantiation ((tycos, vs, sort), params)) |
5caea089dd17
more structural sharing between common target Generic_Target.init
haftmann
parents:
66335
diff
changeset
|
712 |
#> fold (Variable.declare_typ o TFree) vs |
5caea089dd17
more structural sharing between common target Generic_Target.init
haftmann
parents:
66335
diff
changeset
|
713 |
#> fold (Variable.declare_names o Free o snd) params |
5caea089dd17
more structural sharing between common target Generic_Target.init
haftmann
parents:
66335
diff
changeset
|
714 |
#> (Overloading.map_improvable_syntax) (K {primary_constraints = primary_constraints, |
5caea089dd17
more structural sharing between common target Generic_Target.init
haftmann
parents:
66335
diff
changeset
|
715 |
secondary_constraints = [], improve = improve, subst = K NONE, |
5caea089dd17
more structural sharing between common target Generic_Target.init
haftmann
parents:
66335
diff
changeset
|
716 |
no_subst_in_abbrev_mode = false, unchecks = []}) |
5caea089dd17
more structural sharing between common target Generic_Target.init
haftmann
parents:
66335
diff
changeset
|
717 |
#> Overloading.activate_improvable_syntax |
5caea089dd17
more structural sharing between common target Generic_Target.init
haftmann
parents:
66335
diff
changeset
|
718 |
#> Context.proof_map (Syntax_Phases.term_check 0 "resorting" resort_check) |
5caea089dd17
more structural sharing between common target Generic_Target.init
haftmann
parents:
66335
diff
changeset
|
719 |
#> synchronize_inst_syntax, |
5caea089dd17
more structural sharing between common target Generic_Target.init
haftmann
parents:
66335
diff
changeset
|
720 |
conclude = conclude} |
38382 | 721 |
{define = Generic_Target.define foundation, |
60341
fcbd7f0c52c3
clearly separated target primitives (target_foo) from self-contained target operations (foo)
haftmann
parents:
60338
diff
changeset
|
722 |
notes = Generic_Target.notes Generic_Target.theory_target_notes, |
fcbd7f0c52c3
clearly separated target primitives (target_foo) from self-contained target operations (foo)
haftmann
parents:
60338
diff
changeset
|
723 |
abbrev = Generic_Target.abbrev Generic_Target.theory_target_abbrev, |
47279
4bab649dedf0
clarified standard_declaration vs. theory_declaration;
wenzelm
parents:
47250
diff
changeset
|
724 |
declaration = K Generic_Target.theory_declaration, |
73793
26c0ccf17f31
more accurate export morphism enables proper instantiation by interpretation
haftmann
parents:
72516
diff
changeset
|
725 |
theory_registration = registration (Proof_Context.init_global thy), |
61890
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
haftmann
parents:
61262
diff
changeset
|
726 |
locale_dependency = fn _ => error "Not possible in instantiation target", |
66335
a849ce33923d
treat exit separate from regular local theory operations
haftmann
parents:
63347
diff
changeset
|
727 |
pretty = pretty} |
25485 | 728 |
end; |
729 |
||
38382 | 730 |
fun instantiation_cmd arities thy = |
731 |
instantiation (read_multi_arity thy arities) thy; |
|
26238 | 732 |
|
25485 | 733 |
fun gen_instantiation_instance do_proof after_qed lthy = |
734 |
let |
|
25864 | 735 |
val (tycos, vs, sort) = (#arities o the_instantiation) lthy; |
736 |
val arities_proof = maps (fn tyco => Logic.mk_arities (tyco, map snd vs, sort)) tycos; |
|
25462 | 737 |
fun after_qed' results = |
51685
385ef6706252
more standard module name Axclass (according to file name);
wenzelm
parents:
51578
diff
changeset
|
738 |
Local_Theory.background_theory (fold (Axclass.add_arity o Thm.varifyT_global) results) |
25462 | 739 |
#> after_qed; |
740 |
in |
|
741 |
lthy |
|
742 |
|> do_proof after_qed' arities_proof |
|
743 |
end; |
|
744 |
||
25485 | 745 |
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
|
746 |
Proof.theorem NONE (after_qed o map the_single) (map (fn t => [(t, [])]) ts)); |
25462 | 747 |
|
25485 | 748 |
fun prove_instantiation_instance tac = gen_instantiation_instance (fn after_qed => |
25502 | 749 |
fn ts => fn lthy => after_qed (map (fn t => Goal.prove lthy [] [] t |
750 |
(fn {context, ...} => tac context)) ts) lthy) I; |
|
25462 | 751 |
|
28666 | 752 |
fun prove_instantiation_exit tac = prove_instantiation_instance tac |
33671 | 753 |
#> Local_Theory.exit_global; |
28666 | 754 |
|
755 |
fun prove_instantiation_exit_result f tac x lthy = |
|
756 |
let |
|
42360 | 757 |
val morph = Proof_Context.export_morphism lthy |
758 |
(Proof_Context.init_global (Proof_Context.theory_of lthy)); |
|
29439 | 759 |
val y = f morph x; |
28666 | 760 |
in |
761 |
lthy |
|
762 |
|> prove_instantiation_exit (fn ctxt => tac ctxt y) |
|
763 |
|> pair y |
|
764 |
end; |
|
765 |
||
69829 | 766 |
fun theory_map_result arities f g tac = |
767 |
instantiation arities |
|
768 |
#> g |
|
769 |
#-> prove_instantiation_exit_result f tac; |
|
770 |
||
29526 | 771 |
|
31635 | 772 |
(* simplified instantiation interface with no class parameter *) |
773 |
||
31869 | 774 |
fun instance_arity_cmd raw_arities thy = |
31635 | 775 |
let |
31869 | 776 |
val (tycos, vs, sort) = read_multi_arity thy raw_arities; |
777 |
val sorts = map snd vs; |
|
778 |
val arities = maps (fn tyco => Logic.mk_arities (tyco, sorts, sort)) tycos; |
|
47078 | 779 |
fun after_qed results = |
51685
385ef6706252
more standard module name Axclass (according to file name);
wenzelm
parents:
51578
diff
changeset
|
780 |
Proof_Context.background_theory ((fold o fold) Axclass.add_arity results); |
31635 | 781 |
in |
782 |
thy |
|
42360 | 783 |
|> Proof_Context.init_global |
36323
655e2d74de3a
modernized naming conventions of main Isar proof elements;
wenzelm
parents:
36106
diff
changeset
|
784 |
|> Proof.theorem NONE after_qed (map (fn t => [(t, [])]) arities) |
31635 | 785 |
end; |
786 |
||
787 |
||
60618
4c79543cc376
renamed "default" to "standard", to make semantically clear what it is;
wenzelm
parents:
60347
diff
changeset
|
788 |
|
29526 | 789 |
(** tactics and methods **) |
790 |
||
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
59423
diff
changeset
|
791 |
fun intro_classes_tac ctxt facts st = |
29526 | 792 |
let |
60816 | 793 |
val thy = Proof_Context.theory_of ctxt; |
29526 | 794 |
val classes = Sign.all_classes thy; |
795 |
val class_trivs = map (Thm.class_triv thy) classes; |
|
51685
385ef6706252
more standard module name Axclass (according to file name);
wenzelm
parents:
51578
diff
changeset
|
796 |
val class_intros = map_filter (try (#intro o Axclass.get_info thy)) classes; |
29526 | 797 |
val assm_intros = all_assm_intros thy; |
798 |
in |
|
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
59423
diff
changeset
|
799 |
Method.intros_tac ctxt (class_trivs @ class_intros @ assm_intros) facts st |
29526 | 800 |
end; |
801 |
||
60618
4c79543cc376
renamed "default" to "standard", to make semantically clear what it is;
wenzelm
parents:
60347
diff
changeset
|
802 |
fun standard_intro_classes_tac ctxt facts st = |
4c79543cc376
renamed "default" to "standard", to make semantically clear what it is;
wenzelm
parents:
60347
diff
changeset
|
803 |
if null facts andalso not (Thm.no_prems st) then |
69017 | 804 |
(intro_classes_tac ctxt [] ORELSE |
805 |
Locale.intro_locales_tac {strict = true, eager = true} ctxt []) st |
|
60618
4c79543cc376
renamed "default" to "standard", to make semantically clear what it is;
wenzelm
parents:
60347
diff
changeset
|
806 |
else no_tac st; |
29526 | 807 |
|
60619
7512716f03db
no arguments for "standard" (or old "default") methods;
wenzelm
parents:
60618
diff
changeset
|
808 |
fun standard_tac ctxt facts = |
7512716f03db
no arguments for "standard" (or old "default") methods;
wenzelm
parents:
60618
diff
changeset
|
809 |
HEADGOAL (Method.some_rule_tac ctxt [] facts) ORELSE |
60618
4c79543cc376
renamed "default" to "standard", to make semantically clear what it is;
wenzelm
parents:
60347
diff
changeset
|
810 |
standard_intro_classes_tac ctxt facts; |
4c79543cc376
renamed "default" to "standard", to make semantically clear what it is;
wenzelm
parents:
60347
diff
changeset
|
811 |
|
53171 | 812 |
val _ = Theory.setup |
67147 | 813 |
(Method.setup \<^binding>\<open>intro_classes\<close> (Scan.succeed (METHOD o intro_classes_tac)) |
30515 | 814 |
"back-chain introduction rules of classes" #> |
67147 | 815 |
Method.setup \<^binding>\<open>standard\<close> (Scan.succeed (METHOD o standard_tac)) |
62939 | 816 |
"standard proof step: Pure intro/elim rule or class introduction"); |
29526 | 817 |
|
59296
002d817b4c37
formal pretty bodies for class specifications, accepting additional formal bookkeeping in locale.ML
haftmann
parents:
59150
diff
changeset
|
818 |
|
60249
09377954133b
tuned output to resemble input syntax more closely;
wenzelm
parents:
60246
diff
changeset
|
819 |
|
59296
002d817b4c37
formal pretty bodies for class specifications, accepting additional formal bookkeeping in locale.ML
haftmann
parents:
59150
diff
changeset
|
820 |
(** diagnostics **) |
002d817b4c37
formal pretty bodies for class specifications, accepting additional formal bookkeeping in locale.ML
haftmann
parents:
59150
diff
changeset
|
821 |
|
60249
09377954133b
tuned output to resemble input syntax more closely;
wenzelm
parents:
60246
diff
changeset
|
822 |
fun pretty_specification thy class = |
09377954133b
tuned output to resemble input syntax more closely;
wenzelm
parents:
60246
diff
changeset
|
823 |
if is_class thy class then |
59383
1434ef1e0ede
clarified Class.pretty_specification: imitate input source;
wenzelm
parents:
59296
diff
changeset
|
824 |
let |
60249
09377954133b
tuned output to resemble input syntax more closely;
wenzelm
parents:
60246
diff
changeset
|
825 |
val class_ctxt = init class thy; |
09377954133b
tuned output to resemble input syntax more closely;
wenzelm
parents:
60246
diff
changeset
|
826 |
val prt_class = Name_Space.pretty class_ctxt (Proof_Context.class_space class_ctxt); |
09377954133b
tuned output to resemble input syntax more closely;
wenzelm
parents:
60246
diff
changeset
|
827 |
|
09377954133b
tuned output to resemble input syntax more closely;
wenzelm
parents:
60246
diff
changeset
|
828 |
val super_classes = Sign.minimize_sort thy (Sign.super_classes thy class); |
59383
1434ef1e0ede
clarified Class.pretty_specification: imitate input source;
wenzelm
parents:
59296
diff
changeset
|
829 |
|
1434ef1e0ede
clarified Class.pretty_specification: imitate input source;
wenzelm
parents:
59296
diff
changeset
|
830 |
val fix_args = |
60249
09377954133b
tuned output to resemble input syntax more closely;
wenzelm
parents:
60246
diff
changeset
|
831 |
#params (Axclass.get_info thy class) |
59383
1434ef1e0ede
clarified Class.pretty_specification: imitate input source;
wenzelm
parents:
59296
diff
changeset
|
832 |
|> map (fn (c, T) => (Binding.name (Long_Name.base_name c), SOME T, NoSyn)); |
1434ef1e0ede
clarified Class.pretty_specification: imitate input source;
wenzelm
parents:
59296
diff
changeset
|
833 |
val fixes = if null fix_args then [] else [Element.Fixes fix_args]; |
60249
09377954133b
tuned output to resemble input syntax more closely;
wenzelm
parents:
60246
diff
changeset
|
834 |
val assumes = Locale.hyp_spec_of thy class; |
59383
1434ef1e0ede
clarified Class.pretty_specification: imitate input source;
wenzelm
parents:
59296
diff
changeset
|
835 |
|
1434ef1e0ede
clarified Class.pretty_specification: imitate input source;
wenzelm
parents:
59296
diff
changeset
|
836 |
val header = |
60249
09377954133b
tuned output to resemble input syntax more closely;
wenzelm
parents:
60246
diff
changeset
|
837 |
[Pretty.keyword1 "class", Pretty.brk 1, prt_class class, Pretty.str " =", Pretty.brk 1] @ |
09377954133b
tuned output to resemble input syntax more closely;
wenzelm
parents:
60246
diff
changeset
|
838 |
Pretty.separate " +" (map prt_class super_classes) @ |
09377954133b
tuned output to resemble input syntax more closely;
wenzelm
parents:
60246
diff
changeset
|
839 |
(if null super_classes then [] else [Pretty.str " +"]); |
59383
1434ef1e0ede
clarified Class.pretty_specification: imitate input source;
wenzelm
parents:
59296
diff
changeset
|
840 |
val body = |
59385 | 841 |
if null fixes andalso null assumes then [] |
842 |
else |
|
843 |
maps (Element.pretty_ctxt_no_attribs class_ctxt) (fixes @ assumes) |
|
844 |
|> maps (fn prt => [Pretty.fbrk, prt]); |
|
59423
3a0044e95eba
backed out obsolete workaround from ef1edfb36af7
haftmann
parents:
59420
diff
changeset
|
845 |
in if null body then [] else [Pretty.block (header @ body)] end |
3a0044e95eba
backed out obsolete workaround from ef1edfb36af7
haftmann
parents:
59420
diff
changeset
|
846 |
else []; |
59296
002d817b4c37
formal pretty bodies for class specifications, accepting additional formal bookkeeping in locale.ML
haftmann
parents:
59150
diff
changeset
|
847 |
|
24218 | 848 |
end; |