| author | haftmann | 
| Wed, 19 Mar 2008 07:20:32 +0100 | |
| changeset 26328 | b2d6f520172c | 
| parent 26246 | e212c22f35c2 | 
| child 26516 | 1bf210ac0a90 | 
| permissions | -rw-r--r-- | 
| 
404
 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 
wenzelm 
parents:  
diff
changeset
 | 
1  | 
(* Title: Pure/axclass.ML  | 
| 
 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 
wenzelm 
parents:  
diff
changeset
 | 
2  | 
ID: $Id$  | 
| 
 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 
wenzelm 
parents:  
diff
changeset
 | 
3  | 
Author: Markus Wenzel, TU Muenchen  | 
| 
 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 
wenzelm 
parents:  
diff
changeset
 | 
4  | 
|
| 24964 | 5  | 
Type classes defined as predicates, associated with a record of  | 
6  | 
parameters.  | 
|
| 
404
 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 
wenzelm 
parents:  
diff
changeset
 | 
7  | 
*)  | 
| 
 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 
wenzelm 
parents:  
diff
changeset
 | 
8  | 
|
| 
 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 
wenzelm 
parents:  
diff
changeset
 | 
9  | 
signature AX_CLASS =  | 
| 3938 | 10  | 
sig  | 
| 24589 | 11  | 
val define_class: bstring * class list -> string list ->  | 
12  | 
((bstring * attribute list) * term list) list -> theory -> class * theory  | 
|
13  | 
val add_classrel: thm -> theory -> theory  | 
|
14  | 
val add_arity: thm -> theory -> theory  | 
|
15  | 
val prove_classrel: class * class -> tactic -> theory -> theory  | 
|
16  | 
val prove_arity: string * sort list * sort -> tactic -> theory -> theory  | 
|
| 24964 | 17  | 
val get_info: theory -> class ->  | 
18  | 
    {def: thm, intro: thm, axioms: thm list, params: (string * typ) list}
 | 
|
| 
19243
 
5dcb899a8486
moved read_class, read/cert_classrel/arity to sign.ML;
 
wenzelm 
parents: 
19134 
diff
changeset
 | 
19  | 
val class_intros: theory -> thm list  | 
| 20107 | 20  | 
val class_of_param: theory -> string -> class option  | 
| 19405 | 21  | 
val cert_classrel: theory -> class * class -> class * class  | 
22  | 
val read_classrel: theory -> xstring * xstring -> class * class  | 
|
| 
24929
 
408becab067e
renamed AxClass.get_definition to AxClass.get_info (again);
 
wenzelm 
parents: 
24920 
diff
changeset
 | 
23  | 
val axiomatize_class: bstring * class list -> theory -> theory  | 
| 
 
408becab067e
renamed AxClass.get_definition to AxClass.get_info (again);
 
wenzelm 
parents: 
24920 
diff
changeset
 | 
24  | 
val axiomatize_class_cmd: bstring * xstring list -> theory -> theory  | 
| 
 
408becab067e
renamed AxClass.get_definition to AxClass.get_info (again);
 
wenzelm 
parents: 
24920 
diff
changeset
 | 
25  | 
val axiomatize_classrel: (class * class) list -> theory -> theory  | 
| 
 
408becab067e
renamed AxClass.get_definition to AxClass.get_info (again);
 
wenzelm 
parents: 
24920 
diff
changeset
 | 
26  | 
val axiomatize_classrel_cmd: (xstring * xstring) list -> theory -> theory  | 
| 
 
408becab067e
renamed AxClass.get_definition to AxClass.get_info (again);
 
wenzelm 
parents: 
24920 
diff
changeset
 | 
27  | 
val axiomatize_arity: arity -> theory -> theory  | 
| 
 
408becab067e
renamed AxClass.get_definition to AxClass.get_info (again);
 
wenzelm 
parents: 
24920 
diff
changeset
 | 
28  | 
val axiomatize_arity_cmd: xstring * string list * string -> theory -> theory  | 
| 25486 | 29  | 
val instance_name: string * class -> string  | 
| 
25597
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
30  | 
val declare_overloaded: string * typ -> theory -> term * theory  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
31  | 
val define_overloaded: string -> string * term -> theory -> thm * theory  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
32  | 
val inst_tyco_of: theory -> string * typ -> string option  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
33  | 
val unoverload: theory -> thm -> thm  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
34  | 
val overload: theory -> thm -> thm  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
35  | 
val unoverload_conv: theory -> conv  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
36  | 
val overload_conv: theory -> conv  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
37  | 
val unoverload_const: theory -> string * typ -> string  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
38  | 
val param_of_inst: theory -> string * string -> string  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
39  | 
val inst_of_param: theory -> string -> (string * string) option  | 
| 19574 | 40  | 
type cache  | 
| 
24929
 
408becab067e
renamed AxClass.get_definition to AxClass.get_info (again);
 
wenzelm 
parents: 
24920 
diff
changeset
 | 
41  | 
val of_sort: theory -> typ * sort -> cache -> thm list * cache (*exception Sorts.CLASS_ERROR*)  | 
| 19574 | 42  | 
val cache: cache  | 
| 26246 | 43  | 
val introN: string  | 
| 25617 | 44  | 
val axiomsN: string  | 
| 3938 | 45  | 
end;  | 
| 
404
 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 
wenzelm 
parents:  
diff
changeset
 | 
46  | 
|
| 15801 | 47  | 
structure AxClass: AX_CLASS =  | 
| 
404
 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 
wenzelm 
parents:  
diff
changeset
 | 
48  | 
struct  | 
| 
 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 
wenzelm 
parents:  
diff
changeset
 | 
49  | 
|
| 19405 | 50  | 
(** theory data **)  | 
| 423 | 51  | 
|
| 19405 | 52  | 
(* class parameters (canonical order) *)  | 
| 423 | 53  | 
|
| 19405 | 54  | 
type param = string * class;  | 
| 423 | 55  | 
|
| 19405 | 56  | 
fun add_param pp ((x, c): param) params =  | 
57  | 
(case AList.lookup (op =) params x of  | 
|
58  | 
NONE => (x, c) :: params  | 
|
59  | 
  | SOME c' => error ("Duplicate class parameter " ^ quote x ^
 | 
|
60  | 
" for " ^ Pretty.string_of_sort pp [c] ^  | 
|
61  | 
(if c = c' then "" else " and " ^ Pretty.string_of_sort pp [c'])));  | 
|
| 423 | 62  | 
|
| 19405 | 63  | 
fun merge_params _ ([], qs) = qs  | 
64  | 
| merge_params pp (ps, qs) =  | 
|
65  | 
fold_rev (fn q => if member (op =) ps q then I else add_param pp q) qs ps;  | 
|
| 423 | 66  | 
|
67  | 
||
| 19511 | 68  | 
(* axclasses *)  | 
| 6379 | 69  | 
|
| 
19243
 
5dcb899a8486
moved read_class, read/cert_classrel/arity to sign.ML;
 
wenzelm 
parents: 
19134 
diff
changeset
 | 
70  | 
val introN = "intro";  | 
| 19511 | 71  | 
val superN = "super";  | 
| 
19243
 
5dcb899a8486
moved read_class, read/cert_classrel/arity to sign.ML;
 
wenzelm 
parents: 
19134 
diff
changeset
 | 
72  | 
val axiomsN = "axioms";  | 
| 
 
5dcb899a8486
moved read_class, read/cert_classrel/arity to sign.ML;
 
wenzelm 
parents: 
19134 
diff
changeset
 | 
73  | 
|
| 19392 | 74  | 
datatype axclass = AxClass of  | 
| 19460 | 75  | 
 {def: thm,
 | 
| 19392 | 76  | 
intro: thm,  | 
| 
21463
 
42dd50268c8b
completed class parameter handling in axclass.ML
 
haftmann 
parents: 
20628 
diff
changeset
 | 
77  | 
axioms: thm list,  | 
| 21925 | 78  | 
params: (string * typ) list};  | 
| 19392 | 79  | 
|
| 19460 | 80  | 
type axclasses = axclass Symtab.table * param list;  | 
| 19392 | 81  | 
|
| 21925 | 82  | 
fun make_axclass ((def, intro, axioms), params) =  | 
83  | 
  AxClass {def = def, intro = intro, axioms = axioms, params = params};
 | 
|
| 19405 | 84  | 
|
85  | 
fun merge_axclasses pp ((tab1, params1), (tab2, params2)) : axclasses =  | 
|
86  | 
(Symtab.merge (K true) (tab1, tab2), merge_params pp (params1, params2));  | 
|
87  | 
||
| 19392 | 88  | 
|
89  | 
(* instances *)  | 
|
90  | 
||
| 20628 | 91  | 
val classrel_prefix = "classrel_";  | 
92  | 
val arity_prefix = "arity_";  | 
|
| 19511 | 93  | 
|
| 19574 | 94  | 
type instances =  | 
95  | 
((class * class) * thm) list *  | 
|
96  | 
((class * sort list) * thm) list Symtab.table;  | 
|
| 6379 | 97  | 
|
| 19574 | 98  | 
fun merge_instances ((classrel1, arities1): instances, (classrel2, arities2)) =  | 
99  | 
(merge (eq_fst op =) (classrel1, classrel2),  | 
|
100  | 
Symtab.join (K (merge (eq_fst op =))) (arities1, arities2));  | 
|
| 19392 | 101  | 
|
102  | 
||
| 
25597
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
103  | 
(* instance parameters *)  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
104  | 
|
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
105  | 
type inst_params =  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
106  | 
(string * thm) Symtab.table Symtab.table  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
107  | 
(*constant name ~> type constructor ~> (constant name, equation)*)  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
108  | 
* (string * string) Symtab.table; (*constant name ~> (constant name, type constructor)*)  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
109  | 
|
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
110  | 
fun merge_inst_params ((const_param1, param_const1), (const_param2, param_const2)) =  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
111  | 
(Symtab.join (K (Symtab.merge (K true))) (const_param1, const_param2),  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
112  | 
Symtab.merge (K true) (param_const1, param_const2));  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
113  | 
|
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
114  | 
|
| 19511 | 115  | 
(* setup data *)  | 
| 19392 | 116  | 
|
117  | 
structure AxClassData = TheoryDataFun  | 
|
| 22846 | 118  | 
(  | 
| 
25597
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
119  | 
type T = axclasses * (instances * inst_params);  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
120  | 
val empty = ((Symtab.empty, []), (([], Symtab.empty), (Symtab.empty, Symtab.empty)));  | 
| 19574 | 121  | 
val copy = I;  | 
122  | 
val extend = I;  | 
|
| 
25597
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
123  | 
fun merge pp ((axclasses1, (instances1, inst_params1)), (axclasses2, (instances2, inst_params2))) =  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
124  | 
(merge_axclasses pp (axclasses1, axclasses2),  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
125  | 
(merge_instances (instances1, instances2), merge_inst_params (inst_params1, inst_params2)));  | 
| 22846 | 126  | 
);  | 
| 6379 | 127  | 
|
128  | 
||
| 19574 | 129  | 
(* maintain axclasses *)  | 
| 19392 | 130  | 
|
| 19574 | 131  | 
val get_axclasses = #1 o AxClassData.get;  | 
| 
25597
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
132  | 
val map_axclasses = AxClassData.map o apfst;  | 
| 19574 | 133  | 
|
134  | 
val lookup_def = Symtab.lookup o #1 o get_axclasses;  | 
|
| 6379 | 135  | 
|
| 
24929
 
408becab067e
renamed AxClass.get_definition to AxClass.get_info (again);
 
wenzelm 
parents: 
24920 
diff
changeset
 | 
136  | 
fun get_info thy c =  | 
| 19511 | 137  | 
(case lookup_def thy c of  | 
| 19392 | 138  | 
SOME (AxClass info) => info  | 
| 21919 | 139  | 
  | NONE => error ("No such axclass: " ^ quote c));
 | 
| 6379 | 140  | 
|
| 19123 | 141  | 
fun class_intros thy =  | 
| 19392 | 142  | 
let  | 
143  | 
fun add_intro c =  | 
|
| 19511 | 144  | 
      (case lookup_def thy c of SOME (AxClass {intro, ...}) => cons intro | _ => I);
 | 
| 
21931
 
314f9e2a442c
replaced Sign.classes by Sign.all_classes (topologically sorted);
 
wenzelm 
parents: 
21925 
diff
changeset
 | 
145  | 
val classes = Sign.all_classes thy;  | 
| 19392 | 146  | 
in map (Thm.class_triv thy) classes @ fold add_intro classes [] end;  | 
| 
15876
 
a67343c6ab2a
sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
 
wenzelm 
parents: 
15853 
diff
changeset
 | 
147  | 
|
| 
 
a67343c6ab2a
sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
 
wenzelm 
parents: 
15853 
diff
changeset
 | 
148  | 
|
| 19460 | 149  | 
fun get_params thy pred =  | 
| 19574 | 150  | 
let val params = #2 (get_axclasses thy);  | 
| 19460 | 151  | 
in fold (fn (x, c) => if pred c then cons x else I) params [] end;  | 
152  | 
||
153  | 
fun params_of thy c = get_params thy (fn c' => c' = c);  | 
|
154  | 
fun all_params_of thy S = get_params thy (fn c => Sign.subsort thy (S, [c]));  | 
|
155  | 
||
| 24964 | 156  | 
fun class_of_param thy = AList.lookup (op =) (#2 (get_axclasses thy));  | 
| 19460 | 157  | 
|
| 21919 | 158  | 
|
| 19511 | 159  | 
(* maintain instances *)  | 
| 19503 | 160  | 
|
| 25486 | 161  | 
fun instance_name (a, c) = NameSpace.base c ^ "_" ^ NameSpace.base a;  | 
162  | 
||
| 
25597
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
163  | 
val get_instances = #1 o #2 o AxClassData.get;  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
164  | 
val map_instances = AxClassData.map o apsnd o apfst;  | 
| 19528 | 165  | 
|
166  | 
||
| 19574 | 167  | 
fun the_classrel thy (c1, c2) =  | 
168  | 
(case AList.lookup (op =) (#1 (get_instances thy)) (c1, c2) of  | 
|
169  | 
SOME th => Thm.transfer thy th  | 
|
| 24920 | 170  | 
  | NONE => error ("Unproven class relation " ^
 | 
171  | 
Syntax.string_of_classrel (ProofContext.init thy) [c1, c2]));  | 
|
| 19574 | 172  | 
|
173  | 
fun put_classrel arg = map_instances (fn (classrel, arities) =>  | 
|
174  | 
(insert (eq_fst op =) arg classrel, arities));  | 
|
| 19503 | 175  | 
|
176  | 
||
| 19574 | 177  | 
fun the_arity thy a (c, Ss) =  | 
| 
24929
 
408becab067e
renamed AxClass.get_definition to AxClass.get_info (again);
 
wenzelm 
parents: 
24920 
diff
changeset
 | 
178  | 
(case AList.lookup (op =) (Symtab.lookup_list (#2 (get_instances thy)) a) (c, Ss) of  | 
| 19574 | 179  | 
SOME th => Thm.transfer thy th  | 
| 24920 | 180  | 
  | NONE => error ("Unproven type arity " ^
 | 
181  | 
Syntax.string_of_arity (ProofContext.init thy) (a, Ss, [c])));  | 
|
| 19503 | 182  | 
|
| 19574 | 183  | 
fun put_arity ((t, Ss, c), th) = map_instances (fn (classrel, arities) =>  | 
184  | 
(classrel, arities |> Symtab.insert_list (eq_fst op =) (t, ((c, Ss), th))));  | 
|
| 19503 | 185  | 
|
186  | 
||
| 
25597
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
187  | 
(* maintain instance parameters *)  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
188  | 
|
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
189  | 
val get_inst_params = #2 o #2 o AxClassData.get;  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
190  | 
val map_inst_params = AxClassData.map o apsnd o apsnd;  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
191  | 
|
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
192  | 
fun get_inst_param thy (c, tyco) =  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
193  | 
(the o Symtab.lookup ((the o Symtab.lookup (fst (get_inst_params thy))) c)) tyco;  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
194  | 
|
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
195  | 
fun add_inst_param (c, tyco) inst = (map_inst_params o apfst  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
196  | 
o Symtab.map_default (c, Symtab.empty)) (Symtab.update_new (tyco, inst))  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
197  | 
#> (map_inst_params o apsnd) (Symtab.update_new (fst inst, (c, tyco)));  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
198  | 
|
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
199  | 
val inst_of_param = Symtab.lookup o snd o get_inst_params;  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
200  | 
val param_of_inst = fst oo get_inst_param;  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
201  | 
|
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
202  | 
fun inst_thms thy = (Symtab.fold (Symtab.fold (cons o snd o snd) o snd) o fst)  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
203  | 
(get_inst_params thy) [];  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
204  | 
|
| 25863 | 205  | 
fun inst_tyco_of thy = try (fst o dest_Type o the_single o Sign.const_typargs thy);  | 
| 
25597
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
206  | 
|
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
207  | 
fun unoverload thy = MetaSimplifier.simplify true (inst_thms thy);  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
208  | 
fun overload thy = MetaSimplifier.simplify true (map Thm.symmetric (inst_thms thy));  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
209  | 
|
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
210  | 
fun unoverload_conv thy = MetaSimplifier.rewrite true (inst_thms thy);  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
211  | 
fun overload_conv thy = MetaSimplifier.rewrite true (map Thm.symmetric (inst_thms thy));  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
212  | 
|
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
213  | 
fun unoverload_const thy (c_ty as (c, _)) =  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
214  | 
case class_of_param thy c  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
215  | 
of SOME class => (case inst_tyco_of thy c_ty  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
216  | 
of SOME tyco => try (param_of_inst thy) (c, tyco) |> the_default c  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
217  | 
| NONE => c)  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
218  | 
| NONE => c;  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
219  | 
|
| 
404
 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 
wenzelm 
parents:  
diff
changeset
 | 
220  | 
|
| 19511 | 221  | 
(** instances **)  | 
| 
19418
 
03b01c9314fc
reworded add_axclass(_i): canonical specification format,
 
wenzelm 
parents: 
19405 
diff
changeset
 | 
222  | 
|
| 
 
03b01c9314fc
reworded add_axclass(_i): canonical specification format,
 
wenzelm 
parents: 
19405 
diff
changeset
 | 
223  | 
(* class relations *)  | 
| 
 
03b01c9314fc
reworded add_axclass(_i): canonical specification format,
 
wenzelm 
parents: 
19405 
diff
changeset
 | 
224  | 
|
| 19405 | 225  | 
fun cert_classrel thy raw_rel =  | 
| 
15876
 
a67343c6ab2a
sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
 
wenzelm 
parents: 
15853 
diff
changeset
 | 
226  | 
let  | 
| 19405 | 227  | 
val (c1, c2) = pairself (Sign.certify_class thy) raw_rel;  | 
| 24271 | 228  | 
val _ = Sign.primitive_classrel (c1, c2) (Theory.copy thy);  | 
| 19405 | 229  | 
val _ =  | 
| 19460 | 230  | 
(case subtract (op =) (all_params_of thy [c1]) (all_params_of thy [c2]) of  | 
| 19405 | 231  | 
[] => ()  | 
232  | 
      | xs => raise TYPE ("Class " ^ Sign.string_of_sort thy [c1] ^ " lacks parameter(s) " ^
 | 
|
233  | 
commas_quote xs ^ " of " ^ Sign.string_of_sort thy [c2], [], []));  | 
|
234  | 
in (c1, c2) end;  | 
|
235  | 
||
236  | 
fun read_classrel thy raw_rel =  | 
|
237  | 
cert_classrel thy (pairself (Sign.read_class thy) raw_rel)  | 
|
238  | 
handle TYPE (msg, _, _) => error msg;  | 
|
239  | 
||
240  | 
||
241  | 
(* primitive rules *)  | 
|
242  | 
||
243  | 
fun add_classrel th thy =  | 
|
244  | 
let  | 
|
245  | 
    fun err () = raise THM ("add_classrel: malformed class relation", 0, [th]);
 | 
|
| 22691 | 246  | 
val prop = Thm.plain_prop_of (Thm.transfer thy th);  | 
| 19405 | 247  | 
val rel = Logic.dest_classrel prop handle TERM _ => err ();  | 
248  | 
val (c1, c2) = cert_classrel thy rel handle TYPE _ => err ();  | 
|
| 19574 | 249  | 
in  | 
250  | 
thy  | 
|
251  | 
|> Sign.primitive_classrel (c1, c2)  | 
|
252  | 
|> put_classrel ((c1, c2), Drule.unconstrainTs th)  | 
|
253  | 
end;  | 
|
| 
15876
 
a67343c6ab2a
sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
 
wenzelm 
parents: 
15853 
diff
changeset
 | 
254  | 
|
| 19405 | 255  | 
fun add_arity th thy =  | 
| 
15876
 
a67343c6ab2a
sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
 
wenzelm 
parents: 
15853 
diff
changeset
 | 
256  | 
let  | 
| 19528 | 257  | 
    fun err () = raise THM ("add_arity: malformed type arity", 0, [th]);
 | 
| 22691 | 258  | 
val prop = Thm.plain_prop_of (Thm.transfer thy th);  | 
| 19528 | 259  | 
val (t, Ss, c) = Logic.dest_arity prop handle TERM _ => err ();  | 
| 24731 | 260  | 
val _ = map (Sign.certify_sort thy) Ss = Ss orelse err ();  | 
| 
25597
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
261  | 
(*FIXME turn this into a mere check as soon as "attach" has disappeared*)  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
262  | 
val missing_params = Sign.complete_sort thy [c]  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
263  | 
      |> maps (these o Option.map (fn AxClass { params, ... } => params) o lookup_def thy)
 | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
264  | 
|> filter_out (fn (p, _) => can (get_inst_param thy) (p, t));  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
265  | 
fun declare_missing (p, T0) thy =  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
266  | 
let  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
267  | 
val name_inst = instance_name (t, c) ^ "_inst";  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
268  | 
val p' = NameSpace.base p ^ "_" ^ NameSpace.base t;  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
269  | 
val vs = Name.names Name.context Name.aT (replicate (Sign.arity_number thy t) []);  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
270  | 
val T = map_atyps (fn _ => Type (t, map TFree vs)) T0;  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
271  | 
in  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
272  | 
thy  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
273  | 
|> Sign.sticky_prefix name_inst  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
274  | 
|> Sign.no_base_names  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
275  | 
|> Sign.declare_const [] (p', T, NoSyn)  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
276  | 
|-> (fn const' as Const (p'', _) => Thm.add_def false true  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
277  | 
(Thm.def_name p', Logic.mk_equals (const', Const (p, T)))  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
278  | 
#>> Thm.varifyT  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
279  | 
#-> (fn thm => add_inst_param (p, t) (p'', Thm.symmetric thm)  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
280  | 
#> PureThy.note Thm.internalK (p', thm)  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
281  | 
#> snd  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
282  | 
#> Sign.restore_naming thy))  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
283  | 
end;  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
284  | 
|
| 19574 | 285  | 
in  | 
286  | 
thy  | 
|
287  | 
|> Sign.primitive_arity (t, Ss, [c])  | 
|
288  | 
|> put_arity ((t, Ss, c), Drule.unconstrainTs th)  | 
|
| 
25597
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
289  | 
|> fold declare_missing missing_params  | 
| 19574 | 290  | 
end;  | 
| 
404
 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 
wenzelm 
parents:  
diff
changeset
 | 
291  | 
|
| 
 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 
wenzelm 
parents:  
diff
changeset
 | 
292  | 
|
| 
19243
 
5dcb899a8486
moved read_class, read/cert_classrel/arity to sign.ML;
 
wenzelm 
parents: 
19134 
diff
changeset
 | 
293  | 
(* tactical proofs *)  | 
| 
15876
 
a67343c6ab2a
sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
 
wenzelm 
parents: 
15853 
diff
changeset
 | 
294  | 
|
| 19405 | 295  | 
fun prove_classrel raw_rel tac thy =  | 
| 
15876
 
a67343c6ab2a
sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
 
wenzelm 
parents: 
15853 
diff
changeset
 | 
296  | 
let  | 
| 24920 | 297  | 
val ctxt = ProofContext.init thy;  | 
| 19405 | 298  | 
val (c1, c2) = cert_classrel thy raw_rel;  | 
| 24920 | 299  | 
val th = Goal.prove ctxt [] []  | 
| 20049 | 300  | 
(Logic.mk_classrel (c1, c2)) (fn _ => tac) handle ERROR msg =>  | 
| 
19243
 
5dcb899a8486
moved read_class, read/cert_classrel/arity to sign.ML;
 
wenzelm 
parents: 
19134 
diff
changeset
 | 
301  | 
      cat_error msg ("The error(s) above occurred while trying to prove class relation " ^
 | 
| 24920 | 302  | 
quote (Syntax.string_of_classrel ctxt [c1, c2]));  | 
| 19511 | 303  | 
in  | 
304  | 
thy  | 
|
| 20628 | 305  | 
|> PureThy.add_thms [((prefix classrel_prefix (Logic.name_classrel (c1, c2)), th), [])]  | 
| 19511 | 306  | 
|-> (fn [th'] => add_classrel th')  | 
307  | 
end;  | 
|
| 
404
 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 
wenzelm 
parents:  
diff
changeset
 | 
308  | 
|
| 
19243
 
5dcb899a8486
moved read_class, read/cert_classrel/arity to sign.ML;
 
wenzelm 
parents: 
19134 
diff
changeset
 | 
309  | 
fun prove_arity raw_arity tac thy =  | 
| 
15876
 
a67343c6ab2a
sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
 
wenzelm 
parents: 
15853 
diff
changeset
 | 
310  | 
let  | 
| 24920 | 311  | 
val ctxt = ProofContext.init thy;  | 
| 
19243
 
5dcb899a8486
moved read_class, read/cert_classrel/arity to sign.ML;
 
wenzelm 
parents: 
19134 
diff
changeset
 | 
312  | 
val arity = Sign.cert_arity thy raw_arity;  | 
| 20628 | 313  | 
val names = map (prefix arity_prefix) (Logic.name_arities arity);  | 
| 19405 | 314  | 
val props = Logic.mk_arities arity;  | 
| 24920 | 315  | 
val ths = Goal.prove_multi ctxt [] [] props  | 
| 21687 | 316  | 
(fn _ => Goal.precise_conjunction_tac (length props) 1 THEN tac) handle ERROR msg =>  | 
| 
19243
 
5dcb899a8486
moved read_class, read/cert_classrel/arity to sign.ML;
 
wenzelm 
parents: 
19134 
diff
changeset
 | 
317  | 
        cat_error msg ("The error(s) above occurred while trying to prove type arity " ^
 | 
| 24920 | 318  | 
quote (Syntax.string_of_arity ctxt arity));  | 
| 19511 | 319  | 
in  | 
320  | 
thy  | 
|
| 20628 | 321  | 
|> PureThy.add_thms (map (rpair []) (names ~~ ths))  | 
| 19511 | 322  | 
|-> fold add_arity  | 
323  | 
end;  | 
|
324  | 
||
325  | 
||
| 
25597
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
326  | 
(* instance parameters and overloaded definitions *)  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
327  | 
|
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
328  | 
(* declaration and definition of instances of overloaded constants *)  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
329  | 
|
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
330  | 
fun declare_overloaded (c, T) thy =  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
331  | 
let  | 
| 25605 | 332  | 
val class = case class_of_param thy c  | 
333  | 
of SOME class => class  | 
|
334  | 
      | NONE => error ("Not a class parameter: " ^ quote c);
 | 
|
335  | 
val tyco = case inst_tyco_of thy (c, T)  | 
|
336  | 
of SOME tyco => tyco  | 
|
337  | 
      | NONE => error ("Illegal type for instantiation of class parameter: "
 | 
|
338  | 
^ quote (c ^ " :: " ^ Sign.string_of_typ thy T));  | 
|
| 
25597
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
339  | 
val name_inst = instance_name (tyco, class) ^ "_inst";  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
340  | 
val c' = NameSpace.base c ^ "_" ^ NameSpace.base tyco;  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
341  | 
val T' = Type.strip_sorts T;  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
342  | 
in  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
343  | 
thy  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
344  | 
|> Sign.sticky_prefix name_inst  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
345  | 
|> Sign.no_base_names  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
346  | 
|> Sign.declare_const [] (c', T', NoSyn)  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
347  | 
|-> (fn const' as Const (c'', _) => Thm.add_def false true  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
348  | 
(Thm.def_name c', Logic.mk_equals (Const (c, T'), const'))  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
349  | 
#>> Thm.varifyT  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
350  | 
#-> (fn thm => add_inst_param (c, tyco) (c'', thm)  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
351  | 
#> PureThy.note Thm.internalK (c', thm)  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
352  | 
#> snd  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
353  | 
#> Sign.restore_naming thy  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
354  | 
#> pair (Const (c, T))))  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
355  | 
end;  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
356  | 
|
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
357  | 
fun define_overloaded name (c, t) thy =  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
358  | 
let  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
359  | 
val T = Term.fastype_of t;  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
360  | 
val SOME tyco = inst_tyco_of thy (c, T);  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
361  | 
val (c', eq) = get_inst_param thy (c, tyco);  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
362  | 
val prop = Logic.mk_equals (Const (c', T), t);  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
363  | 
val name' = Thm.def_name_optional  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
364  | 
(NameSpace.base c ^ "_" ^ NameSpace.base tyco) name;  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
365  | 
in  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
366  | 
thy  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
367  | 
|> Thm.add_def false false (name', prop)  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
368  | 
|>> (fn thm => Drule.transitive_thm OF [eq, thm])  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
369  | 
end;  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
370  | 
|
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
371  | 
|
| 19511 | 372  | 
|
373  | 
(** class definitions **)  | 
|
374  | 
||
| 23421 | 375  | 
fun split_defined n eq =  | 
376  | 
let  | 
|
377  | 
val intro =  | 
|
378  | 
(eq RS Drule.equal_elim_rule2)  | 
|
379  | 
|> Conjunction.curry_balanced n  | 
|
380  | 
|> n = 0 ? Thm.eq_assumption 1;  | 
|
381  | 
val dests =  | 
|
382  | 
if n = 0 then []  | 
|
383  | 
else  | 
|
384  | 
(eq RS Drule.equal_elim_rule1)  | 
|
385  | 
|> BalancedTree.dest (fn th =>  | 
|
386  | 
(th RS Conjunction.conjunctionD1, th RS Conjunction.conjunctionD2)) n;  | 
|
387  | 
in (intro, dests) end;  | 
|
388  | 
||
| 24964 | 389  | 
fun define_class (bclass, raw_super) raw_params raw_specs thy =  | 
| 19511 | 390  | 
let  | 
391  | 
val ctxt = ProofContext.init thy;  | 
|
| 24964 | 392  | 
val pp = Syntax.pp ctxt;  | 
| 19511 | 393  | 
|
394  | 
||
| 24964 | 395  | 
(* class *)  | 
| 19511 | 396  | 
|
397  | 
val bconst = Logic.const_of_class bclass;  | 
|
398  | 
val class = Sign.full_name thy bclass;  | 
|
| 24731 | 399  | 
val super = Sign.minimize_sort thy (Sign.certify_sort thy raw_super);  | 
| 19511 | 400  | 
|
| 24964 | 401  | 
fun check_constraint (a, S) =  | 
402  | 
if Sign.subsort thy (super, S) then ()  | 
|
403  | 
      else error ("Sort constraint of type variable " ^
 | 
|
404  | 
setmp show_sorts true (Pretty.string_of_typ pp) (TFree (a, S)) ^  | 
|
405  | 
" needs to be weaker than " ^ Pretty.string_of_sort pp super);  | 
|
406  | 
||
407  | 
||
408  | 
(* params *)  | 
|
409  | 
||
410  | 
val params = raw_params |> map (fn p =>  | 
|
411  | 
let  | 
|
412  | 
val T = Sign.the_const_type thy p;  | 
|
413  | 
val _ =  | 
|
414  | 
(case Term.add_tvarsT T [] of  | 
|
415  | 
[((a, _), S)] => check_constraint (a, S)  | 
|
416  | 
          | _ => error ("Exactly one type variable expected in class parameter " ^ quote p));
 | 
|
417  | 
val T' = Term.map_type_tvar (fn _ => TFree (Name.aT, [class])) T;  | 
|
418  | 
in (p, T') end);  | 
|
419  | 
||
420  | 
||
421  | 
(* axioms *)  | 
|
422  | 
||
| 19511 | 423  | 
fun prep_axiom t =  | 
424  | 
(case Term.add_tfrees t [] of  | 
|
| 24964 | 425  | 
[(a, S)] => check_constraint (a, S)  | 
426  | 
| [] => ()  | 
|
427  | 
      | _ => error ("Multiple type variables in class axiom:\n" ^ Pretty.string_of_term pp t);
 | 
|
428  | 
t  | 
|
429  | 
|> Term.map_types (Term.map_atyps (fn TFree _ => Term.aT [] | U => U))  | 
|
430  | 
|> Logic.close_form);  | 
|
| 19511 | 431  | 
|
| 24681 | 432  | 
val axiomss = map (map (prep_axiom o Sign.cert_prop thy) o snd) raw_specs;  | 
| 22745 | 433  | 
val name_atts = map fst raw_specs;  | 
| 19511 | 434  | 
|
435  | 
||
436  | 
(* definition *)  | 
|
437  | 
||
438  | 
val conjs = map (curry Logic.mk_inclass (Term.aT [])) super @ flat axiomss;  | 
|
439  | 
val class_eq =  | 
|
| 23421 | 440  | 
Logic.mk_equals (Logic.mk_inclass (Term.aT [], class), Logic.mk_conjunction_balanced conjs);  | 
| 19511 | 441  | 
|
442  | 
val ([def], def_thy) =  | 
|
443  | 
thy  | 
|
444  | 
|> Sign.primitive_class (bclass, super)  | 
|
445  | 
|> PureThy.add_defs_i false [((Thm.def_name bconst, class_eq), [])];  | 
|
446  | 
val (raw_intro, (raw_classrel, raw_axioms)) =  | 
|
| 23421 | 447  | 
split_defined (length conjs) def ||> chop (length super);  | 
| 19392 | 448  | 
|
| 
19418
 
03b01c9314fc
reworded add_axclass(_i): canonical specification format,
 
wenzelm 
parents: 
19405 
diff
changeset
 | 
449  | 
|
| 19511 | 450  | 
(* facts *)  | 
451  | 
||
452  | 
val class_triv = Thm.class_triv def_thy class;  | 
|
453  | 
val ([(_, [intro]), (_, classrel), (_, axioms)], facts_thy) =  | 
|
454  | 
def_thy  | 
|
455  | 
|> PureThy.note_thmss_qualified "" bconst  | 
|
456  | 
[((introN, []), [([Drule.standard raw_intro], [])]),  | 
|
457  | 
((superN, []), [(map Drule.standard raw_classrel, [])]),  | 
|
458  | 
((axiomsN, []), [(map (fn th => Drule.standard (class_triv RS th)) raw_axioms, [])])];  | 
|
459  | 
||
| 24847 | 460  | 
|
| 19511 | 461  | 
(* result *)  | 
462  | 
||
| 24964 | 463  | 
val axclass = make_axclass ((def, intro, axioms), params);  | 
| 19511 | 464  | 
val result_thy =  | 
465  | 
facts_thy  | 
|
| 19574 | 466  | 
|> fold put_classrel (map (pair class) super ~~ classrel)  | 
| 19511 | 467  | 
|> Sign.add_path bconst  | 
468  | 
|> PureThy.note_thmss_i "" (name_atts ~~ map Thm.simple_fact (unflat axiomss axioms)) |> snd  | 
|
469  | 
|> Sign.restore_naming facts_thy  | 
|
| 19574 | 470  | 
|> map_axclasses (fn (axclasses, parameters) =>  | 
| 21919 | 471  | 
(Symtab.update (class, axclass) axclasses,  | 
| 24964 | 472  | 
fold (fn (x, _) => add_param pp (x, class)) params parameters));  | 
| 19511 | 473  | 
|
474  | 
in (class, result_thy) end;  | 
|
475  | 
||
476  | 
||
| 19531 | 477  | 
|
| 19511 | 478  | 
(** axiomatizations **)  | 
479  | 
||
480  | 
local  | 
|
481  | 
||
| 20628 | 482  | 
fun axiomatize prep mk name add raw_args thy =  | 
483  | 
let  | 
|
484  | 
val args = prep thy raw_args;  | 
|
485  | 
val specs = mk args;  | 
|
486  | 
val names = name args;  | 
|
487  | 
in thy |> PureThy.add_axioms_i (map (rpair []) (names ~~ specs)) |-> fold add end;  | 
|
| 19511 | 488  | 
|
489  | 
fun ax_classrel prep =  | 
|
| 20628 | 490  | 
axiomatize (map o prep) (map Logic.mk_classrel)  | 
491  | 
(map (prefix classrel_prefix o Logic.name_classrel)) add_classrel;  | 
|
| 19511 | 492  | 
|
493  | 
fun ax_arity prep =  | 
|
| 20628 | 494  | 
axiomatize prep Logic.mk_arities  | 
495  | 
(map (prefix arity_prefix) o Logic.name_arities) add_arity;  | 
|
| 19511 | 496  | 
|
| 19706 | 497  | 
fun class_const c =  | 
498  | 
(Logic.const_of_class c, Term.itselfT (Term.aT []) --> propT);  | 
|
499  | 
||
| 19511 | 500  | 
fun ax_class prep_class prep_classrel (bclass, raw_super) thy =  | 
501  | 
let  | 
|
502  | 
val class = Sign.full_name thy bclass;  | 
|
| 24731 | 503  | 
val super = map (prep_class thy) raw_super |> Sign.minimize_sort thy;  | 
| 19511 | 504  | 
in  | 
505  | 
thy  | 
|
506  | 
|> Sign.primitive_class (bclass, super)  | 
|
507  | 
|> ax_classrel prep_classrel (map (fn c => (class, c)) super)  | 
|
| 19706 | 508  | 
|> Theory.add_deps "" (class_const class) (map class_const super)  | 
| 19511 | 509  | 
end;  | 
510  | 
||
511  | 
in  | 
|
512  | 
||
| 
24929
 
408becab067e
renamed AxClass.get_definition to AxClass.get_info (again);
 
wenzelm 
parents: 
24920 
diff
changeset
 | 
513  | 
val axiomatize_class = ax_class Sign.certify_class cert_classrel;  | 
| 
 
408becab067e
renamed AxClass.get_definition to AxClass.get_info (again);
 
wenzelm 
parents: 
24920 
diff
changeset
 | 
514  | 
val axiomatize_class_cmd = ax_class Sign.read_class read_classrel;  | 
| 
 
408becab067e
renamed AxClass.get_definition to AxClass.get_info (again);
 
wenzelm 
parents: 
24920 
diff
changeset
 | 
515  | 
val axiomatize_classrel = ax_classrel cert_classrel;  | 
| 
 
408becab067e
renamed AxClass.get_definition to AxClass.get_info (again);
 
wenzelm 
parents: 
24920 
diff
changeset
 | 
516  | 
val axiomatize_classrel_cmd = ax_classrel read_classrel;  | 
| 
 
408becab067e
renamed AxClass.get_definition to AxClass.get_info (again);
 
wenzelm 
parents: 
24920 
diff
changeset
 | 
517  | 
val axiomatize_arity = ax_arity Sign.cert_arity;  | 
| 
 
408becab067e
renamed AxClass.get_definition to AxClass.get_info (again);
 
wenzelm 
parents: 
24920 
diff
changeset
 | 
518  | 
val axiomatize_arity_cmd = ax_arity Sign.read_arity;  | 
| 19511 | 519  | 
|
520  | 
end;  | 
|
521  | 
||
522  | 
||
523  | 
||
524  | 
(** explicit derivations -- cached **)  | 
|
525  | 
||
| 19574 | 526  | 
datatype cache = Types of (class * thm) list Typtab.table;  | 
527  | 
||
| 19511 | 528  | 
local  | 
| 19503 | 529  | 
|
| 19574 | 530  | 
fun lookup_type (Types cache) = AList.lookup (op =) o Typtab.lookup_list cache;  | 
531  | 
fun insert_type T der (Types cache) = Types (Typtab.insert_list (eq_fst op =) (T, der) cache);  | 
|
| 19503 | 532  | 
|
| 19522 | 533  | 
fun derive_type _ (_, []) = []  | 
534  | 
| derive_type thy (typ, sort) =  | 
|
535  | 
let  | 
|
| 19528 | 536  | 
val vars = Term.fold_atyps  | 
| 19522 | 537  | 
(fn T as TFree (_, S) => insert (eq_fst op =) (T, S)  | 
538  | 
| T as TVar (_, S) => insert (eq_fst op =) (T, S)  | 
|
| 19528 | 539  | 
| _ => I) typ [];  | 
| 19574 | 540  | 
val hyps = vars  | 
541  | 
|> map (fn (T, S) => (T, Drule.sort_triv thy (T, S) ~~ S));  | 
|
| 19642 | 542  | 
val ths = (typ, sort) |> Sorts.of_sort_derivation (Sign.pp thy) (Sign.classes_of thy)  | 
| 
22570
 
f166a5416b3f
renamed of_sort_derivation record fields (avoid clash with Alice keywords);
 
wenzelm 
parents: 
22385 
diff
changeset
 | 
543  | 
           {class_relation =
 | 
| 19574 | 544  | 
fn (th, c1) => fn c2 => th RS the_classrel thy (c1, c2),  | 
| 
22570
 
f166a5416b3f
renamed of_sort_derivation record fields (avoid clash with Alice keywords);
 
wenzelm 
parents: 
22385 
diff
changeset
 | 
545  | 
type_constructor =  | 
| 19574 | 546  | 
fn a => fn dom => fn c =>  | 
547  | 
let val Ss = map (map snd) dom and ths = maps (map fst) dom  | 
|
548  | 
in ths MRS the_arity thy a (c, Ss) end,  | 
|
| 
22570
 
f166a5416b3f
renamed of_sort_derivation record fields (avoid clash with Alice keywords);
 
wenzelm 
parents: 
22385 
diff
changeset
 | 
549  | 
type_variable =  | 
| 19574 | 550  | 
the_default [] o AList.lookup (op =) hyps};  | 
551  | 
in ths end;  | 
|
| 19503 | 552  | 
|
| 19511 | 553  | 
in  | 
554  | 
||
| 19574 | 555  | 
fun of_sort thy (typ, sort) cache =  | 
556  | 
let  | 
|
557  | 
val sort' = filter (is_none o lookup_type cache typ) sort;  | 
|
558  | 
val ths' = derive_type thy (typ, sort')  | 
|
559  | 
      handle ERROR msg => cat_error msg ("The error(s) above occurred for sort derivation: " ^
 | 
|
560  | 
Sign.string_of_typ thy typ ^ " :: " ^ Sign.string_of_sort thy sort');  | 
|
561  | 
val cache' = cache |> fold (insert_type typ) (sort' ~~ ths');  | 
|
562  | 
val ths =  | 
|
563  | 
sort |> map (fn c =>  | 
|
564  | 
Goal.finish (the (lookup_type cache' typ c) RS  | 
|
565  | 
Goal.init (Thm.cterm_of thy (Logic.mk_inclass (typ, c))))  | 
|
| 20260 | 566  | 
|> Thm.adjust_maxidx_thm ~1);  | 
| 19574 | 567  | 
in (ths, cache') end;  | 
| 19503 | 568  | 
|
| 19511 | 569  | 
end;  | 
| 
19418
 
03b01c9314fc
reworded add_axclass(_i): canonical specification format,
 
wenzelm 
parents: 
19405 
diff
changeset
 | 
570  | 
|
| 
24929
 
408becab067e
renamed AxClass.get_definition to AxClass.get_info (again);
 
wenzelm 
parents: 
24920 
diff
changeset
 | 
571  | 
val cache = Types Typtab.empty;  | 
| 
 
408becab067e
renamed AxClass.get_definition to AxClass.get_info (again);
 
wenzelm 
parents: 
24920 
diff
changeset
 | 
572  | 
|
| 
15876
 
a67343c6ab2a
sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
 
wenzelm 
parents: 
15853 
diff
changeset
 | 
573  | 
end;  |