| author | haftmann | 
| Tue, 31 Aug 2010 14:21:06 +0200 | |
| changeset 38926 | 24f82786cc57 | 
| parent 38757 | 2b3e054ae6fc | 
| child 39133 | 70d3915c92f0 | 
| 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  | 
| 29526 | 17  | 
val print_classes: theory -> 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,  | 
|
| 
28715
 
238f9966c80e
class morphism stemming from locale interpretation
 
haftmann 
parents: 
28674 
diff
changeset
 | 
72  | 
|
| 
 
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  | 
|
| 24218 | 152  | 
fun print_classes thy =  | 
153  | 
let  | 
|
| 
36610
 
bafd82950e24
renamed ProofContext.init to ProofContext.init_global to emphasize that this is not the real thing;
 
wenzelm 
parents: 
36462 
diff
changeset
 | 
154  | 
val ctxt = ProofContext.init_global thy;  | 
| 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;  | 
| 24218 | 166  | 
fun mk_param (c, ty) = Pretty.str (Sign.extern_const thy c ^ " :: "  | 
| 
32966
 
5b21661fe618
indicate CRITICAL nature of various setmp combinators;
 
wenzelm 
parents: 
32850 
diff
changeset
 | 
167  | 
^ setmp_CRITICAL show_sorts false (Syntax.string_of_typ ctxt o Type.strip_sorts) ty);  | 
| 24218 | 168  | 
fun mk_entry class = (Pretty.block o Pretty.fbreaks o map_filter I) [  | 
| 25062 | 169  | 
      (SOME o Pretty.str) ("class " ^ Sign.extern_class thy class ^ ":"),
 | 
| 24218 | 170  | 
(SOME o Pretty.block) [Pretty.str "supersort: ",  | 
| 24920 | 171  | 
(Syntax.pretty_sort ctxt o Sign.minimize_sort thy o Sign.super_classes thy) class],  | 
| 25062 | 172  | 
((fn [] => NONE | ps => (SOME o Pretty.block o Pretty.fbreaks)  | 
173  | 
(Pretty.str "parameters:" :: ps)) o map mk_param  | 
|
| 
24930
 
cc2e0e8c81af
renamed AxClass.get_definition to AxClass.get_info (again);
 
wenzelm 
parents: 
24920 
diff
changeset
 | 
174  | 
o these o Option.map #params o try (AxClass.get_info thy)) class,  | 
| 24218 | 175  | 
(SOME o Pretty.block o Pretty.breaks) [  | 
176  | 
Pretty.str "instances:",  | 
|
177  | 
Pretty.list "" "" (map (mk_arity class) (the_arities class))  | 
|
178  | 
]  | 
|
179  | 
]  | 
|
180  | 
in  | 
|
| 24589 | 181  | 
(Pretty.writeln o Pretty.chunks o separate (Pretty.str "")  | 
182  | 
o map mk_entry o Sorts.all_classes) algebra  | 
|
| 24218 | 183  | 
end;  | 
184  | 
||
185  | 
||
186  | 
(* updaters *)  | 
|
187  | 
||
| 32850 | 188  | 
fun register class sups params base_sort base_morph export_morph  | 
| 29358 | 189  | 
axiom assm_intro of_class thy =  | 
| 25002 | 190  | 
let  | 
| 
25368
 
f12613fda79d
proper implementation of check phase; non-qualified names for class operations
 
haftmann 
parents: 
25344 
diff
changeset
 | 
191  | 
val operations = map (fn (v_ty as (_, ty), (c, _)) =>  | 
| 25683 | 192  | 
(c, (class, (ty, Free v_ty)))) params;  | 
| 25038 | 193  | 
val add_class = Graph.new_node (class,  | 
| 31599 | 194  | 
make_class_data (((map o pairself) fst params, base_sort,  | 
| 32850 | 195  | 
base_morph, export_morph, assm_intro, of_class, axiom), ([], operations)))  | 
| 29526 | 196  | 
#> fold (curry Graph.add_edge class) sups;  | 
| 25618 | 197  | 
in ClassData.map add_class thy end;  | 
| 24218 | 198  | 
|
| 
36674
 
d95f39448121
eq_morphism is always optional: avoid trivial morphism for empty list of equations
 
haftmann 
parents: 
36672 
diff
changeset
 | 
199  | 
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
 | 
200  | 
of SOME eq_morph => fold (fn cls => fn thy =>  | 
| 38107 | 201  | 
Context.theory_map (Locale.amend_registration (cls, base_morphism thy cls)  | 
202  | 
(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
 | 
203  | 
| NONE => thy;  | 
| 29526 | 204  | 
|
| 
25368
 
f12613fda79d
proper implementation of check phase; non-qualified names for class operations
 
haftmann 
parents: 
25344 
diff
changeset
 | 
205  | 
fun register_operation class (c, (t, some_def)) thy =  | 
| 25062 | 206  | 
let  | 
| 29358 | 207  | 
val base_sort = base_sort thy class;  | 
| 29439 | 208  | 
val prep_typ = map_type_tfree  | 
209  | 
(fn (v, sort) => if Name.aT = v  | 
|
210  | 
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
 | 
211  | 
val t' = map_types prep_typ t;  | 
| 
 
f12613fda79d
proper implementation of check phase; non-qualified names for class operations
 
haftmann 
parents: 
25344 
diff
changeset
 | 
212  | 
val ty' = Term.fastype_of t';  | 
| 25062 | 213  | 
in  | 
214  | 
thy  | 
|
215  | 
|> (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
 | 
216  | 
(fn (defs, operations) =>  | 
| 25096 | 217  | 
(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
 | 
218  | 
(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
 | 
219  | 
|> activate_defs class (the_list some_def)  | 
| 25062 | 220  | 
end;  | 
| 24218 | 221  | 
|
| 29558 | 222  | 
fun register_subclass (sub, sup) some_dep_morph some_wit export thy =  | 
| 25618 | 223  | 
let  | 
| 29558 | 224  | 
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
 | 
225  | 
[Option.map (Drule.export_without_context_open o Element.conclude_witness) some_wit,  | 
| 29558 | 226  | 
(fst o rules thy) sub];  | 
| 29610 | 227  | 
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
 | 
228  | 
val classrel = Skip_Proof.prove_global thy [] [] (Logic.mk_classrel (sub, sup))  | 
| 29558 | 229  | 
(K tac);  | 
| 
29509
 
1ff0f3f08a7b
migrated class package to new locale implementation
 
haftmann 
parents: 
29439 
diff
changeset
 | 
230  | 
val diff_sort = Sign.complete_sort thy [sup]  | 
| 31012 | 231  | 
|> subtract (op =) (Sign.complete_sort thy [sub])  | 
232  | 
|> filter (is_class thy);  | 
|
| 32074 | 233  | 
val add_dependency = case some_dep_morph  | 
234  | 
of SOME dep_morph => Locale.add_dependency sub  | 
|
235  | 
(sup, dep_morph $> Element.satisfy_morphism (the_list some_wit)) export  | 
|
236  | 
| NONE => I  | 
|
| 25618 | 237  | 
in  | 
238  | 
thy  | 
|
| 37246 | 239  | 
|> AxClass.add_classrel classrel  | 
| 25618 | 240  | 
|> ClassData.map (Graph.add_edge (sub, sup))  | 
| 29526 | 241  | 
|> activate_defs sub (these_defs thy diff_sort)  | 
| 32074 | 242  | 
|> add_dependency  | 
| 24218 | 243  | 
end;  | 
244  | 
||
245  | 
||
| 24589 | 246  | 
(** classes and class target **)  | 
| 24218 | 247  | 
|
| 25002 | 248  | 
(* class context syntax *)  | 
| 24748 | 249  | 
|
| 
35858
 
0d394a82337e
handle hidden polymorphism in class target (without class target syntax, though)
 
haftmann 
parents: 
35845 
diff
changeset
 | 
250  | 
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
 | 
251  | 
o these_operations thy;  | 
| 29577 | 252  | 
|
| 
29632
 
c3d576157244
fixed reading of class specs: declare class operations in context
 
haftmann 
parents: 
29610 
diff
changeset
 | 
253  | 
fun redeclare_const thy c =  | 
| 
30364
 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 
wenzelm 
parents: 
30344 
diff
changeset
 | 
254  | 
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
 | 
255  | 
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
 | 
256  | 
|
| 29577 | 257  | 
fun synchronize_class_syntax sort base_sort ctxt =  | 
| 
24914
 
95cda5dd58d5
added proper subclass concept; improved class target
 
haftmann 
parents: 
24901 
diff
changeset
 | 
258  | 
let  | 
| 
25344
 
00c2179db769
synchronize_syntax: improved declare_const (still inactive);
 
wenzelm 
parents: 
25326 
diff
changeset
 | 
259  | 
val thy = ProofContext.theory_of ctxt;  | 
| 26596 | 260  | 
val algebra = Sign.classes_of thy;  | 
| 29577 | 261  | 
val operations = these_operations thy sort;  | 
| 26518 | 262  | 
fun subst_class_typ sort = map_type_tfree (K (TVar ((Name.aT, 0), sort)));  | 
263  | 
val primary_constraints =  | 
|
| 
25368
 
f12613fda79d
proper implementation of check phase; non-qualified names for class operations
 
haftmann 
parents: 
25344 
diff
changeset
 | 
264  | 
(map o apsnd) (subst_class_typ base_sort o fst o snd) operations;  | 
| 26518 | 265  | 
val secondary_constraints =  | 
| 
25368
 
f12613fda79d
proper implementation of check phase; non-qualified names for class operations
 
haftmann 
parents: 
25344 
diff
changeset
 | 
266  | 
(map o apsnd) (fn (class, (ty, _)) => subst_class_typ [class] ty) operations;  | 
| 26518 | 267  | 
fun improve (c, ty) = (case AList.lookup (op =) primary_constraints c  | 
| 26238 | 268  | 
of SOME ty' => (case try (Type.raw_match (ty', ty)) Vartab.empty  | 
269  | 
of SOME tyenv => (case Vartab.lookup tyenv (Name.aT, 0)  | 
|
| 33969 | 270  | 
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
 | 
271  | 
if Type_Infer.is_param vi  | 
| 26596 | 272  | 
andalso Sorts.sort_le algebra (base_sort, sort)  | 
273  | 
then SOME (ty', TFree (Name.aT, base_sort))  | 
|
274  | 
else NONE  | 
|
| 26238 | 275  | 
| _ => NONE)  | 
276  | 
| NONE => NONE)  | 
|
277  | 
| NONE => NONE)  | 
|
278  | 
fun subst (c, ty) = Option.map snd (AList.lookup (op =) operations c);  | 
|
| 29577 | 279  | 
val unchecks = these_unchecks thy sort;  | 
| 25083 | 280  | 
in  | 
281  | 
ctxt  | 
|
| 
29632
 
c3d576157244
fixed reading of class specs: declare class operations in context
 
haftmann 
parents: 
29610 
diff
changeset
 | 
282  | 
|> fold (redeclare_const thy o fst) primary_constraints  | 
| 26518 | 283  | 
|> Overloading.map_improvable_syntax (K (((primary_constraints, secondary_constraints),  | 
| 26730 | 284  | 
(((improve, subst), true), unchecks)), false))  | 
| 26518 | 285  | 
|> Overloading.set_primary_constraints  | 
| 25083 | 286  | 
end;  | 
287  | 
||
| 
29632
 
c3d576157244
fixed reading of class specs: declare class operations in context
 
haftmann 
parents: 
29610 
diff
changeset
 | 
288  | 
fun redeclare_operations thy sort =  | 
| 
 
c3d576157244
fixed reading of class specs: declare class operations in context
 
haftmann 
parents: 
29610 
diff
changeset
 | 
289  | 
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
 | 
290  | 
|
| 29577 | 291  | 
fun begin sort base_sort ctxt =  | 
| 25083 | 292  | 
ctxt  | 
293  | 
|> Variable.declare_term  | 
|
294  | 
(Logic.mk_type (TFree (Name.aT, base_sort)))  | 
|
| 29577 | 295  | 
|> synchronize_class_syntax sort base_sort  | 
| 26238 | 296  | 
|> Overloading.add_improvable_syntax;  | 
| 
24901
 
d3cbf79769b9
added first version of user-space type system for class target
 
haftmann 
parents: 
24847 
diff
changeset
 | 
297  | 
|
| 25311 | 298  | 
fun init class thy =  | 
299  | 
thy  | 
|
| 
29509
 
1ff0f3f08a7b
migrated class package to new locale implementation
 
haftmann 
parents: 
29439 
diff
changeset
 | 
300  | 
|> Locale.init class  | 
| 29358 | 301  | 
|> begin [class] (base_sort thy class);  | 
| 
24914
 
95cda5dd58d5
added proper subclass concept; improved class target
 
haftmann 
parents: 
24901 
diff
changeset
 | 
302  | 
|
| 24748 | 303  | 
|
| 
27690
 
24738db98d34
some steps towards explicit class target for canonical interpretation
 
haftmann 
parents: 
27684 
diff
changeset
 | 
304  | 
(* class target *)  | 
| 
 
24738db98d34
some steps towards explicit class target for canonical interpretation
 
haftmann 
parents: 
27684 
diff
changeset
 | 
305  | 
|
| 
30364
 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 
wenzelm 
parents: 
30344 
diff
changeset
 | 
306  | 
val class_prefix = Logic.const_of_class o Long_Name.base_name;  | 
| 29526 | 307  | 
|
| 
38619
 
25e401d53900
tuned: less formal noise in Named_Target, more coherence in Class
 
haftmann 
parents: 
38392 
diff
changeset
 | 
308  | 
fun target_extension f class lthy =  | 
| 
 
25e401d53900
tuned: less formal noise in Named_Target, more coherence in Class
 
haftmann 
parents: 
38392 
diff
changeset
 | 
309  | 
lthy  | 
| 
 
25e401d53900
tuned: less formal noise in Named_Target, more coherence in Class
 
haftmann 
parents: 
38392 
diff
changeset
 | 
310  | 
|> Local_Theory.raw_theory f  | 
| 
 
25e401d53900
tuned: less formal noise in Named_Target, more coherence in Class
 
haftmann 
parents: 
38392 
diff
changeset
 | 
311  | 
|> Local_Theory.target (synchronize_class_syntax [class]  | 
| 
 
25e401d53900
tuned: less formal noise in Named_Target, more coherence in Class
 
haftmann 
parents: 
38392 
diff
changeset
 | 
312  | 
(base_sort (ProofContext.theory_of lthy) class));  | 
| 
 
25e401d53900
tuned: less formal noise in Named_Target, more coherence in Class
 
haftmann 
parents: 
38392 
diff
changeset
 | 
313  | 
|
| 
 
25e401d53900
tuned: less formal noise in Named_Target, more coherence in Class
 
haftmann 
parents: 
38392 
diff
changeset
 | 
314  | 
local  | 
| 
 
25e401d53900
tuned: less formal noise in Named_Target, more coherence in Class
 
haftmann 
parents: 
38392 
diff
changeset
 | 
315  | 
|
| 
 
25e401d53900
tuned: less formal noise in Named_Target, more coherence in Class
 
haftmann 
parents: 
38392 
diff
changeset
 | 
316  | 
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
 | 
317  | 
let  | 
| 29577 | 318  | 
val morph = morphism thy class;  | 
| 
30344
 
10a67c5ddddb
more uniform handling of binding in targets and derived elements;
 
wenzelm 
parents: 
30280 
diff
changeset
 | 
319  | 
val b = Morphism.binding morph c;  | 
| 
 
10a67c5ddddb
more uniform handling of binding in targets and derived elements;
 
wenzelm 
parents: 
30280 
diff
changeset
 | 
320  | 
val b_def = Morphism.binding morph (Binding.suffix_name "_dict" b);  | 
| 29577 | 321  | 
val c' = Sign.full_name thy b;  | 
| 29439 | 322  | 
val dict' = Morphism.term morph dict;  | 
| 
35858
 
0d394a82337e
handle hidden polymorphism in class target (without class target syntax, though)
 
haftmann 
parents: 
35845 
diff
changeset
 | 
323  | 
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
 | 
324  | 
val def_eq = Logic.mk_equals (list_comb (Const (c', ty'), type_params), dict')  | 
| 29577 | 325  | 
|> map_types Type.strip_sorts;  | 
| 
27690
 
24738db98d34
some steps towards explicit class target for canonical interpretation
 
haftmann 
parents: 
27684 
diff
changeset
 | 
326  | 
in  | 
| 29577 | 327  | 
thy  | 
| 
33173
 
b8ca12f6681a
eliminated obsolete tags for types/consts -- now handled via name space, in strongly typed fashion;
 
wenzelm 
parents: 
32970 
diff
changeset
 | 
328  | 
|> Sign.declare_const ((b, Type.strip_sorts ty'), mx)  | 
| 29577 | 329  | 
|> snd  | 
330  | 
|> Thm.add_def false false (b_def, def_eq)  | 
|
| 
36106
 
19deea200358
Thm.add_axiom/add_def: return internal name of foundational axiom;
 
wenzelm 
parents: 
35858 
diff
changeset
 | 
331  | 
|>> apsnd Thm.varifyT_global  | 
| 
 
19deea200358
Thm.add_axiom/add_def: return internal name of foundational axiom;
 
wenzelm 
parents: 
35858 
diff
changeset
 | 
332  | 
|-> (fn (_, def_thm) => PureThy.store_thm (b_def, def_thm)  | 
| 29577 | 333  | 
#> snd  | 
| 
35858
 
0d394a82337e
handle hidden polymorphism in class target (without class target syntax, though)
 
haftmann 
parents: 
35845 
diff
changeset
 | 
334  | 
#> 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
 | 
335  | 
|> Sign.add_const_constraint (c', SOME ty')  | 
| 
 
24738db98d34
some steps towards explicit class target for canonical interpretation
 
haftmann 
parents: 
27684 
diff
changeset
 | 
336  | 
end;  | 
| 
 
24738db98d34
some steps towards explicit class target for canonical interpretation
 
haftmann 
parents: 
27684 
diff
changeset
 | 
337  | 
|
| 
38619
 
25e401d53900
tuned: less formal noise in Named_Target, more coherence in Class
 
haftmann 
parents: 
38392 
diff
changeset
 | 
338  | 
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
 | 
339  | 
let  | 
| 29577 | 340  | 
val morph = morphism thy class;  | 
341  | 
val unchecks = these_unchecks thy [class];  | 
|
| 
30344
 
10a67c5ddddb
more uniform handling of binding in targets and derived elements;
 
wenzelm 
parents: 
30280 
diff
changeset
 | 
342  | 
val b = Morphism.binding morph c;  | 
| 29577 | 343  | 
val c' = Sign.full_name thy b;  | 
| 
27690
 
24738db98d34
some steps towards explicit class target for canonical interpretation
 
haftmann 
parents: 
27684 
diff
changeset
 | 
344  | 
val rhs' = Pattern.rewrite_term thy unchecks [] rhs;  | 
| 
 
24738db98d34
some steps towards explicit class target for canonical interpretation
 
haftmann 
parents: 
27684 
diff
changeset
 | 
345  | 
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
 | 
346  | 
val rhs'' = map_types Logic.varifyT_global rhs';  | 
| 
27690
 
24738db98d34
some steps towards explicit class target for canonical interpretation
 
haftmann 
parents: 
27684 
diff
changeset
 | 
347  | 
in  | 
| 29577 | 348  | 
thy  | 
| 
33173
 
b8ca12f6681a
eliminated obsolete tags for types/consts -- now handled via name space, in strongly typed fashion;
 
wenzelm 
parents: 
32970 
diff
changeset
 | 
349  | 
|> Sign.add_abbrev (#1 prmode) (b, rhs'')  | 
| 29577 | 350  | 
|> snd  | 
| 
27690
 
24738db98d34
some steps towards explicit class target for canonical interpretation
 
haftmann 
parents: 
27684 
diff
changeset
 | 
351  | 
|> Sign.add_const_constraint (c', SOME ty')  | 
| 
 
24738db98d34
some steps towards explicit class target for canonical interpretation
 
haftmann 
parents: 
27684 
diff
changeset
 | 
352  | 
|> 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
 | 
353  | 
|> 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
 | 
354  | 
end;  | 
| 
 
24738db98d34
some steps towards explicit class target for canonical interpretation
 
haftmann 
parents: 
27684 
diff
changeset
 | 
355  | 
|
| 
38619
 
25e401d53900
tuned: less formal noise in Named_Target, more coherence in Class
 
haftmann 
parents: 
38392 
diff
changeset
 | 
356  | 
in  | 
| 
 
25e401d53900
tuned: less formal noise in Named_Target, more coherence in Class
 
haftmann 
parents: 
38392 
diff
changeset
 | 
357  | 
|
| 
 
25e401d53900
tuned: less formal noise in Named_Target, more coherence in Class
 
haftmann 
parents: 
38392 
diff
changeset
 | 
358  | 
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
 | 
359  | 
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
 | 
360  | 
|
| 
 
25e401d53900
tuned: less formal noise in Named_Target, more coherence in Class
 
haftmann 
parents: 
38392 
diff
changeset
 | 
361  | 
end;  | 
| 
 
25e401d53900
tuned: less formal noise in Named_Target, more coherence in Class
 
haftmann 
parents: 
38392 
diff
changeset
 | 
362  | 
|
| 
27690
 
24738db98d34
some steps towards explicit class target for canonical interpretation
 
haftmann 
parents: 
27684 
diff
changeset
 | 
363  | 
|
| 31635 | 364  | 
(* simple subclasses *)  | 
365  | 
||
366  | 
local  | 
|
367  | 
||
368  | 
fun gen_classrel mk_prop classrel thy =  | 
|
369  | 
let  | 
|
370  | 
fun after_qed results =  | 
|
| 
38756
 
d07959fabde6
renamed ProofContext.theory(_result) to ProofContext.background_theory(_result) to emphasize that this belongs to the infrastructure and is rarely appropriate in user-space tools;
 
wenzelm 
parents: 
38619 
diff
changeset
 | 
371  | 
ProofContext.background_theory ((fold o fold) AxClass.add_classrel results);  | 
| 31635 | 372  | 
in  | 
373  | 
thy  | 
|
| 
36610
 
bafd82950e24
renamed ProofContext.init to ProofContext.init_global to emphasize that this is not the real thing;
 
wenzelm 
parents: 
36462 
diff
changeset
 | 
374  | 
|> ProofContext.init_global  | 
| 
36323
 
655e2d74de3a
modernized naming conventions of main Isar proof elements;
 
wenzelm 
parents: 
36106 
diff
changeset
 | 
375  | 
|> Proof.theorem NONE after_qed [[(mk_prop thy classrel, [])]]  | 
| 31635 | 376  | 
end;  | 
377  | 
||
378  | 
in  | 
|
379  | 
||
380  | 
val classrel =  | 
|
381  | 
gen_classrel (Logic.mk_classrel oo AxClass.cert_classrel);  | 
|
382  | 
val classrel_cmd =  | 
|
383  | 
gen_classrel (Logic.mk_classrel oo AxClass.read_classrel);  | 
|
384  | 
||
385  | 
end; (*local*)  | 
|
386  | 
||
387  | 
||
| 25462 | 388  | 
(** instantiation target **)  | 
389  | 
||
390  | 
(* bookkeeping *)  | 
|
391  | 
||
392  | 
datatype instantiation = Instantiation of {
 | 
|
| 25864 | 393  | 
arities: string list * (string * sort) list * sort,  | 
| 25462 | 394  | 
params: ((string * string) * (string * typ)) list  | 
| 25603 | 395  | 
(*(instantiation parameter, type constructor), (local instantiation parameter, typ)*)  | 
| 25462 | 396  | 
}  | 
397  | 
||
| 33519 | 398  | 
structure Instantiation = Proof_Data  | 
| 25462 | 399  | 
(  | 
400  | 
type T = instantiation  | 
|
| 25536 | 401  | 
  fun init _ = Instantiation { arities = ([], [], []), params = [] };
 | 
| 25462 | 402  | 
);  | 
403  | 
||
| 25485 | 404  | 
fun mk_instantiation (arities, params) =  | 
405  | 
  Instantiation { arities = arities, params = params };
 | 
|
| 33671 | 406  | 
fun get_instantiation lthy = case Instantiation.get (Local_Theory.target_of lthy)  | 
| 25485 | 407  | 
of Instantiation data => data;  | 
| 33671 | 408  | 
fun map_instantiation f = (Local_Theory.target o Instantiation.map)  | 
| 25514 | 409  | 
  (fn Instantiation { arities, params } => mk_instantiation (f (arities, params)));
 | 
| 25462 | 410  | 
|
| 25514 | 411  | 
fun the_instantiation lthy = case get_instantiation lthy  | 
| 25536 | 412  | 
 of { arities = ([], [], []), ... } => error "No instantiation target"
 | 
| 25485 | 413  | 
| data => data;  | 
| 25462 | 414  | 
|
| 25485 | 415  | 
val instantiation_params = #params o get_instantiation;  | 
| 25462 | 416  | 
|
| 
30519
 
c05c0199826f
coherent binding policy with primitive target operations
 
haftmann 
parents: 
30364 
diff
changeset
 | 
417  | 
fun instantiation_param lthy b = instantiation_params lthy  | 
| 
 
c05c0199826f
coherent binding policy with primitive target operations
 
haftmann 
parents: 
30364 
diff
changeset
 | 
418  | 
|> find_first (fn (_, (v, _)) => Binding.name_of b = v)  | 
| 25462 | 419  | 
|> Option.map (fst o fst);  | 
420  | 
||
| 31869 | 421  | 
fun read_multi_arity thy (raw_tycos, raw_sorts, raw_sort) =  | 
422  | 
let  | 
|
| 
36610
 
bafd82950e24
renamed ProofContext.init to ProofContext.init_global to emphasize that this is not the real thing;
 
wenzelm 
parents: 
36462 
diff
changeset
 | 
423  | 
val ctxt = ProofContext.init_global thy;  | 
| 
35669
 
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
 
wenzelm 
parents: 
35315 
diff
changeset
 | 
424  | 
val all_arities = map (fn raw_tyco => ProofContext.read_arity ctxt  | 
| 31869 | 425  | 
(raw_tyco, raw_sorts, raw_sort)) raw_tycos;  | 
426  | 
val tycos = map #1 all_arities;  | 
|
427  | 
val (_, sorts, sort) = hd all_arities;  | 
|
428  | 
val vs = Name.names Name.context Name.aT sorts;  | 
|
429  | 
in (tycos, vs, sort) end;  | 
|
430  | 
||
| 25462 | 431  | 
|
432  | 
(* syntax *)  | 
|
433  | 
||
| 26238 | 434  | 
fun synchronize_inst_syntax ctxt =  | 
| 25462 | 435  | 
let  | 
| 33969 | 436  | 
    val Instantiation { params, ... } = Instantiation.get ctxt;
 | 
| 31249 | 437  | 
|
| 33969 | 438  | 
val lookup_inst_param = AxClass.lookup_inst_param  | 
439  | 
(Sign.consts_of (ProofContext.theory_of ctxt)) params;  | 
|
| 31249 | 440  | 
fun subst (c, ty) = case lookup_inst_param (c, ty)  | 
441  | 
of SOME (v_ty as (_, ty)) => SOME (ty, Free v_ty)  | 
|
442  | 
| NONE => NONE;  | 
|
| 26238 | 443  | 
val unchecks =  | 
444  | 
map (fn ((c, _), v_ty as (_, ty)) => (Free v_ty, Const (c, ty))) params;  | 
|
445  | 
in  | 
|
446  | 
ctxt  | 
|
| 26259 | 447  | 
|> Overloading.map_improvable_syntax  | 
| 26730 | 448  | 
(fn (((primary_constraints, _), (((improve, _), _), _)), _) =>  | 
449  | 
(((primary_constraints, []), (((improve, subst), false), unchecks)), false))  | 
|
| 26238 | 450  | 
end;  | 
| 25462 | 451  | 
|
| 38382 | 452  | 
fun resort_terms pp algebra consts constraints ts =  | 
453  | 
let  | 
|
454  | 
fun matchings (Const (c_ty as (c, _))) = (case constraints c  | 
|
455  | 
of NONE => I  | 
|
456  | 
| SOME sorts => fold2 (curry (Sorts.meet_sort algebra))  | 
|
457  | 
(Consts.typargs consts c_ty) sorts)  | 
|
458  | 
| matchings _ = I  | 
|
459  | 
val tvartab = (fold o fold_aterms) matchings ts Vartab.empty  | 
|
460  | 
handle Sorts.CLASS_ERROR e => error (Sorts.class_error pp e);  | 
|
461  | 
val inst = map_type_tvar  | 
|
462  | 
(fn (vi, sort) => TVar (vi, the_default sort (Vartab.lookup tvartab vi)));  | 
|
463  | 
in if Vartab.is_empty tvartab then NONE else SOME ((map o map_types) inst ts) end;  | 
|
464  | 
||
| 25462 | 465  | 
|
466  | 
(* target *)  | 
|
467  | 
||
| 
37678
 
0040bafffdef
"prod" and "sum" replace "*" and "+" respectively
 
haftmann 
parents: 
37393 
diff
changeset
 | 
468  | 
val sanitize_name = (*necessary as long as "dirty" type identifiers are permitted*)  | 
| 25485 | 469  | 
let  | 
| 25574 | 470  | 
fun is_valid s = Symbol.is_ascii_letter s orelse Symbol.is_ascii_digit s  | 
471  | 
orelse s = "'" orelse s = "_";  | 
|
| 25485 | 472  | 
val is_junk = not o is_valid andf Symbol.is_regular;  | 
473  | 
val junk = Scan.many is_junk;  | 
|
474  | 
val scan_valids = Symbol.scanner "Malformed input"  | 
|
475  | 
((junk |--  | 
|
476  | 
(Scan.optional (Scan.one Symbol.is_ascii_letter) "x" ^^ (Scan.many is_valid >> implode)  | 
|
477  | 
--| junk))  | 
|
| 25999 | 478  | 
::: Scan.repeat ((Scan.many1 is_valid >> implode) --| junk));  | 
| 25485 | 479  | 
in  | 
480  | 
explode #> scan_valids #> implode  | 
|
481  | 
end;  | 
|
482  | 
||
| 
37678
 
0040bafffdef
"prod" and "sum" replace "*" and "+" respectively
 
haftmann 
parents: 
37393 
diff
changeset
 | 
483  | 
val type_name = sanitize_name o Long_Name.base_name;  | 
| 26259 | 484  | 
|
| 
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
 | 
485  | 
fun define_overloaded (c, U) v (b_def, rhs) = Local_Theory.background_theory_result  | 
| 38382 | 486  | 
(AxClass.declare_overloaded (c, U) ##>> AxClass.define_overloaded b_def (c, rhs))  | 
487  | 
##> (map_instantiation o apsnd) (filter_out (fn (_, (v', _)) => v' = v))  | 
|
488  | 
##> Local_Theory.target synchronize_inst_syntax;  | 
|
489  | 
||
490  | 
fun foundation (((b, U), mx), (b_def, rhs)) (type_params, term_params) lthy =  | 
|
491  | 
case instantiation_param lthy b  | 
|
492  | 
   of SOME c => if mx <> NoSyn then error ("Illegal mixfix syntax for overloaded constant " ^ quote c)
 | 
|
493  | 
else lthy |> define_overloaded (c, U) (Binding.name_of b) (b_def, rhs)  | 
|
494  | 
| NONE => lthy |>  | 
|
495  | 
Generic_Target.theory_foundation (((b, U), mx), (b_def, rhs)) (type_params, term_params);  | 
|
496  | 
||
497  | 
fun pretty lthy =  | 
|
| 26518 | 498  | 
let  | 
| 38382 | 499  | 
    val { arities = (tycos, vs, sort), params } = the_instantiation lthy;
 | 
500  | 
val thy = ProofContext.theory_of lthy;  | 
|
501  | 
fun pr_arity tyco = Syntax.pretty_arity lthy (tyco, map snd vs, sort);  | 
|
502  | 
fun pr_param ((c, _), (v, ty)) =  | 
|
503  | 
(Pretty.block o Pretty.breaks) [Pretty.str v, Pretty.str "==",  | 
|
504  | 
(Pretty.str o Sign.extern_const thy) c, Pretty.str "::", Syntax.pretty_typ_global thy ty];  | 
|
505  | 
in Pretty.str "instantiation" :: map pr_arity tycos @ map pr_param params end;  | 
|
| 26518 | 506  | 
|
| 38382 | 507  | 
fun conclude lthy =  | 
508  | 
let  | 
|
509  | 
val (tycos, vs, sort) = (#arities o the_instantiation) lthy;  | 
|
510  | 
val thy = ProofContext.theory_of lthy;  | 
|
511  | 
val _ = map (fn tyco => if Sign.of_sort thy  | 
|
512  | 
(Type (tyco, map TFree vs), sort)  | 
|
513  | 
      then () else error ("Missing instance proof for type " ^ quote (Sign.extern_type thy tyco)))
 | 
|
514  | 
tycos;  | 
|
515  | 
in lthy end;  | 
|
516  | 
||
517  | 
fun instantiation (tycos, vs, sort) thy =  | 
|
| 25462 | 518  | 
let  | 
| 25536 | 519  | 
val _ = if null tycos then error "At least one arity must be given" else ();  | 
| 31249 | 520  | 
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
 | 
521  | 
fun get_param tyco (param, (_, (c, ty))) =  | 
| 
 
c3d576157244
fixed reading of class specs: declare class operations in context
 
haftmann 
parents: 
29610 
diff
changeset
 | 
522  | 
if can (AxClass.param_of_inst thy) (c, tyco)  | 
| 25603 | 523  | 
then NONE else SOME ((c, tyco),  | 
| 25864 | 524  | 
(param ^ "_" ^ type_name tyco, map_atyps (K (Type (tyco, map TFree vs))) ty));  | 
| 31249 | 525  | 
val params = map_product get_param tycos class_params |> map_filter I;  | 
| 26518 | 526  | 
val primary_constraints = map (apsnd  | 
| 31249 | 527  | 
(map_atyps (K (TVar ((Name.aT, 0), [])))) o snd o snd) class_params;  | 
| 
26939
 
1035c89b4c02
moved global pretty/string_of functions from Sign to Syntax;
 
wenzelm 
parents: 
26761 
diff
changeset
 | 
528  | 
val pp = Syntax.pp_global thy;  | 
| 26518 | 529  | 
val algebra = Sign.classes_of thy  | 
530  | 
|> fold (fn tyco => Sorts.add_arities pp  | 
|
531  | 
(tyco, map (fn class => (class, map snd vs)) sort)) tycos;  | 
|
532  | 
val consts = Sign.consts_of thy;  | 
|
533  | 
val improve_constraints = AList.lookup (op =)  | 
|
| 31249 | 534  | 
(map (fn (_, (class, (c, _))) => (c, [[class]])) class_params);  | 
| 33967 | 535  | 
fun resort_check ts ctxt = case resort_terms (Syntax.pp ctxt) algebra consts improve_constraints ts  | 
| 26518 | 536  | 
of NONE => NONE  | 
537  | 
| SOME ts' => SOME (ts', ctxt);  | 
|
| 31249 | 538  | 
val lookup_inst_param = AxClass.lookup_inst_param consts params;  | 
| 31210 | 539  | 
val typ_instance = Type.typ_instance (Sign.tsig_of thy);  | 
| 31249 | 540  | 
fun improve (c, ty) = case lookup_inst_param (c, ty)  | 
541  | 
of SOME (_, ty') => if typ_instance (ty', ty) then SOME (ty, ty') else NONE  | 
|
| 26259 | 542  | 
| NONE => NONE;  | 
| 25485 | 543  | 
in  | 
544  | 
thy  | 
|
| 
32379
 
a97e9caebd60
additional checkpoints avoid problems in error situations
 
haftmann 
parents: 
32074 
diff
changeset
 | 
545  | 
|> Theory.checkpoint  | 
| 
36610
 
bafd82950e24
renamed ProofContext.init to ProofContext.init_global to emphasize that this is not the real thing;
 
wenzelm 
parents: 
36462 
diff
changeset
 | 
546  | 
|> ProofContext.init_global  | 
| 31249 | 547  | 
|> Instantiation.put (mk_instantiation ((tycos, vs, sort), params))  | 
| 27281 | 548  | 
|> fold (Variable.declare_typ o TFree) vs  | 
| 31249 | 549  | 
|> fold (Variable.declare_names o Free o snd) params  | 
| 26259 | 550  | 
|> (Overloading.map_improvable_syntax o apfst)  | 
| 
29969
 
9dbb046136d0
more precise improvement in instantiation user space type system
 
haftmann 
parents: 
29632 
diff
changeset
 | 
551  | 
(K ((primary_constraints, []), (((improve, K NONE), false), [])))  | 
| 26259 | 552  | 
|> Overloading.add_improvable_syntax  | 
| 26518 | 553  | 
|> Context.proof_map (Syntax.add_term_check 0 "resorting" resort_check)  | 
| 26238 | 554  | 
|> synchronize_inst_syntax  | 
| 38382 | 555  | 
|> Local_Theory.init NONE ""  | 
556  | 
       {define = Generic_Target.define foundation,
 | 
|
557  | 
notes = Generic_Target.notes  | 
|
558  | 
(fn kind => fn global_facts => fn _ => Generic_Target.theory_notes kind global_facts),  | 
|
559  | 
abbrev = Generic_Target.abbrev  | 
|
560  | 
(fn prmode => fn (b, mx) => fn (t, _) => fn _ => Generic_Target.theory_abbrev prmode ((b, mx), t)),  | 
|
561  | 
declaration = K Generic_Target.theory_declaration,  | 
|
562  | 
syntax_declaration = K Generic_Target.theory_declaration,  | 
|
563  | 
pretty = pretty,  | 
|
564  | 
exit = Local_Theory.target_of o conclude}  | 
|
| 25485 | 565  | 
end;  | 
566  | 
||
| 38382 | 567  | 
fun instantiation_cmd arities thy =  | 
568  | 
instantiation (read_multi_arity thy arities) thy;  | 
|
| 26238 | 569  | 
|
| 25485 | 570  | 
fun gen_instantiation_instance do_proof after_qed lthy =  | 
571  | 
let  | 
|
| 25864 | 572  | 
val (tycos, vs, sort) = (#arities o the_instantiation) lthy;  | 
573  | 
val arities_proof = maps (fn tyco => Logic.mk_arities (tyco, map snd vs, sort)) tycos;  | 
|
| 25462 | 574  | 
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
 | 
575  | 
Local_Theory.background_theory (fold (AxClass.add_arity o Thm.varifyT_global) results)  | 
| 25462 | 576  | 
#> after_qed;  | 
577  | 
in  | 
|
578  | 
lthy  | 
|
579  | 
|> do_proof after_qed' arities_proof  | 
|
580  | 
end;  | 
|
581  | 
||
| 25485 | 582  | 
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
 | 
583  | 
Proof.theorem NONE (after_qed o map the_single) (map (fn t => [(t, [])]) ts));  | 
| 25462 | 584  | 
|
| 25485 | 585  | 
fun prove_instantiation_instance tac = gen_instantiation_instance (fn after_qed =>  | 
| 25502 | 586  | 
fn ts => fn lthy => after_qed (map (fn t => Goal.prove lthy [] [] t  | 
587  | 
    (fn {context, ...} => tac context)) ts) lthy) I;
 | 
|
| 25462 | 588  | 
|
| 28666 | 589  | 
fun prove_instantiation_exit tac = prove_instantiation_instance tac  | 
| 33671 | 590  | 
#> Local_Theory.exit_global;  | 
| 28666 | 591  | 
|
592  | 
fun prove_instantiation_exit_result f tac x lthy =  | 
|
593  | 
let  | 
|
| 29439 | 594  | 
val morph = ProofContext.export_morphism lthy  | 
| 
36610
 
bafd82950e24
renamed ProofContext.init to ProofContext.init_global to emphasize that this is not the real thing;
 
wenzelm 
parents: 
36462 
diff
changeset
 | 
595  | 
(ProofContext.init_global (ProofContext.theory_of lthy));  | 
| 29439 | 596  | 
val y = f morph x;  | 
| 28666 | 597  | 
in  | 
598  | 
lthy  | 
|
599  | 
|> prove_instantiation_exit (fn ctxt => tac ctxt y)  | 
|
600  | 
|> pair y  | 
|
601  | 
end;  | 
|
602  | 
||
| 29526 | 603  | 
|
| 31635 | 604  | 
(* simplified instantiation interface with no class parameter *)  | 
605  | 
||
| 31869 | 606  | 
fun instance_arity_cmd raw_arities thy =  | 
| 31635 | 607  | 
let  | 
| 31869 | 608  | 
val (tycos, vs, sort) = read_multi_arity thy raw_arities;  | 
609  | 
val sorts = map snd vs;  | 
|
610  | 
val arities = maps (fn tyco => Logic.mk_arities (tyco, sorts, sort)) tycos;  | 
|
| 
38756
 
d07959fabde6
renamed ProofContext.theory(_result) to ProofContext.background_theory(_result) to emphasize that this belongs to the infrastructure and is rarely appropriate in user-space tools;
 
wenzelm 
parents: 
38619 
diff
changeset
 | 
611  | 
fun after_qed results = ProofContext.background_theory  | 
| 37246 | 612  | 
((fold o fold) AxClass.add_arity results);  | 
| 31635 | 613  | 
in  | 
614  | 
thy  | 
|
| 
36610
 
bafd82950e24
renamed ProofContext.init to ProofContext.init_global to emphasize that this is not the real thing;
 
wenzelm 
parents: 
36462 
diff
changeset
 | 
615  | 
|> ProofContext.init_global  | 
| 
36323
 
655e2d74de3a
modernized naming conventions of main Isar proof elements;
 
wenzelm 
parents: 
36106 
diff
changeset
 | 
616  | 
|> Proof.theorem NONE after_qed (map (fn t => [(t, [])]) arities)  | 
| 31635 | 617  | 
end;  | 
618  | 
||
619  | 
||
| 29526 | 620  | 
(** tactics and methods **)  | 
621  | 
||
622  | 
fun intro_classes_tac facts st =  | 
|
623  | 
let  | 
|
624  | 
val thy = Thm.theory_of_thm st;  | 
|
625  | 
val classes = Sign.all_classes thy;  | 
|
626  | 
val class_trivs = map (Thm.class_triv thy) classes;  | 
|
627  | 
val class_intros = map_filter (try (#intro o AxClass.get_info thy)) classes;  | 
|
628  | 
val assm_intros = all_assm_intros thy;  | 
|
629  | 
in  | 
|
630  | 
Method.intros_tac (class_trivs @ class_intros @ assm_intros) facts st  | 
|
631  | 
end;  | 
|
632  | 
||
633  | 
fun default_intro_tac ctxt [] =  | 
|
| 29577 | 634  | 
intro_classes_tac [] ORELSE Locale.intro_locales_tac true ctxt []  | 
| 29526 | 635  | 
| default_intro_tac _ _ = no_tac;  | 
636  | 
||
637  | 
fun default_tac rules ctxt facts =  | 
|
638  | 
HEADGOAL (Method.some_rule_tac rules ctxt facts) ORELSE  | 
|
639  | 
default_intro_tac ctxt facts;  | 
|
640  | 
||
641  | 
val _ = Context.>> (Context.map_theory  | 
|
| 30515 | 642  | 
(Method.setup (Binding.name "intro_classes") (Scan.succeed (K (METHOD intro_classes_tac)))  | 
643  | 
"back-chain introduction rules of classes" #>  | 
|
644  | 
Method.setup (Binding.name "default") (Attrib.thms >> (METHOD oo default_tac))  | 
|
645  | 
"apply some intro/elim rule"));  | 
|
| 29526 | 646  | 
|
| 24218 | 647  | 
end;  | 
| 25683 | 648  |