| author | haftmann | 
| Wed, 17 Feb 2010 09:48:53 +0100 | |
| changeset 35160 | 6eb2b6c1d2d5 | 
| parent 35021 | c839a4c670c6 | 
| child 35201 | c2ddb9395436 | 
| 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  | 
Author: Markus Wenzel, TU Muenchen  | 
| 
 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 
wenzelm 
parents:  
diff
changeset
 | 
3  | 
|
| 24964 | 4  | 
Type classes defined as predicates, associated with a record of  | 
5  | 
parameters.  | 
|
| 
404
 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 
wenzelm 
parents:  
diff
changeset
 | 
6  | 
*)  | 
| 
 
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  | 
signature AX_CLASS =  | 
| 3938 | 9  | 
sig  | 
| 
30344
 
10a67c5ddddb
more uniform handling of binding in targets and derived elements;
 
wenzelm 
parents: 
30280 
diff
changeset
 | 
10  | 
val define_class: binding * class list -> string list ->  | 
| 30211 | 11  | 
(Thm.binding * term list) list -> theory -> class * theory  | 
| 24589 | 12  | 
val add_classrel: thm -> theory -> theory  | 
13  | 
val add_arity: thm -> theory -> theory  | 
|
14  | 
val prove_classrel: class * class -> tactic -> theory -> theory  | 
|
15  | 
val prove_arity: string * sort list * sort -> tactic -> theory -> theory  | 
|
| 24964 | 16  | 
val get_info: theory -> class ->  | 
17  | 
    {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
 | 
18  | 
val class_intros: theory -> thm list  | 
| 20107 | 19  | 
val class_of_param: theory -> string -> class option  | 
| 19405 | 20  | 
val cert_classrel: theory -> class * class -> class * class  | 
21  | 
val read_classrel: theory -> xstring * xstring -> class * class  | 
|
| 
30344
 
10a67c5ddddb
more uniform handling of binding in targets and derived elements;
 
wenzelm 
parents: 
30280 
diff
changeset
 | 
22  | 
val axiomatize_class: binding * class list -> theory -> theory  | 
| 
 
10a67c5ddddb
more uniform handling of binding in targets and derived elements;
 
wenzelm 
parents: 
30280 
diff
changeset
 | 
23  | 
val axiomatize_class_cmd: binding * xstring list -> theory -> theory  | 
| 
24929
 
408becab067e
renamed AxClass.get_definition to AxClass.get_info (again);
 
wenzelm 
parents: 
24920 
diff
changeset
 | 
24  | 
val axiomatize_classrel: (class * class) list -> theory -> theory  | 
| 
 
408becab067e
renamed AxClass.get_definition to AxClass.get_info (again);
 
wenzelm 
parents: 
24920 
diff
changeset
 | 
25  | 
val axiomatize_classrel_cmd: (xstring * xstring) list -> theory -> theory  | 
| 
 
408becab067e
renamed AxClass.get_definition to AxClass.get_info (again);
 
wenzelm 
parents: 
24920 
diff
changeset
 | 
26  | 
val axiomatize_arity: arity -> theory -> theory  | 
| 
 
408becab067e
renamed AxClass.get_definition to AxClass.get_info (again);
 
wenzelm 
parents: 
24920 
diff
changeset
 | 
27  | 
val axiomatize_arity_cmd: xstring * string list * string -> theory -> theory  | 
| 25486 | 28  | 
val instance_name: string * class -> string  | 
| 
25597
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
29  | 
val declare_overloaded: string * typ -> theory -> term * theory  | 
| 
30519
 
c05c0199826f
coherent binding policy with primitive target operations
 
haftmann 
parents: 
30469 
diff
changeset
 | 
30  | 
val define_overloaded: binding -> string * term -> theory -> thm * theory  | 
| 
25597
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
31  | 
val unoverload: theory -> thm -> thm  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
32  | 
val overload: theory -> thm -> thm  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
33  | 
val unoverload_conv: theory -> conv  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
34  | 
val overload_conv: theory -> conv  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
35  | 
val unoverload_const: theory -> string * typ -> string  | 
| 31249 | 36  | 
val lookup_inst_param: Consts.T -> ((string * string) * 'a) list -> string * typ -> 'a option  | 
| 
25597
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
37  | 
val param_of_inst: theory -> string * string -> string  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
38  | 
val inst_of_param: theory -> string -> (string * string) option  | 
| 33172 | 39  | 
val thynames_of_arity: theory -> class * string -> string list  | 
| 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  | 
|
| 27397 | 82  | 
fun make_axclass ((def, intro, axioms), params) = AxClass  | 
83  | 
  {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 =  | 
| 33172 | 95  | 
((class * class) * thm) list * (*classrel theorems*)  | 
96  | 
((class * sort list) * (thm * string)) list Symtab.table; (*arity theorems with theory name*)  | 
|
| 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  | 
|
| 
34259
 
2ba492b8b6e8
discontinued old TheoryDataFun, but retain Theory_Data_PP with is Pretty.pp argument to merge (still required in exotic situations -- hard to get rid of);
 
wenzelm 
parents: 
34245 
diff
changeset
 | 
117  | 
structure AxClassData = Theory_Data_PP  | 
| 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 extend = I;  | 
| 
25597
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
122  | 
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
 | 
123  | 
(merge_axclasses pp (axclasses1, axclasses2),  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
124  | 
(merge_instances (instances1, instances2), merge_inst_params (inst_params1, inst_params2)));  | 
| 22846 | 125  | 
);  | 
| 6379 | 126  | 
|
127  | 
||
| 19574 | 128  | 
(* maintain axclasses *)  | 
| 19392 | 129  | 
|
| 19574 | 130  | 
val get_axclasses = #1 o AxClassData.get;  | 
| 
25597
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
131  | 
val map_axclasses = AxClassData.map o apfst;  | 
| 19574 | 132  | 
|
133  | 
val lookup_def = Symtab.lookup o #1 o get_axclasses;  | 
|
| 6379 | 134  | 
|
| 
24929
 
408becab067e
renamed AxClass.get_definition to AxClass.get_info (again);
 
wenzelm 
parents: 
24920 
diff
changeset
 | 
135  | 
fun get_info thy c =  | 
| 19511 | 136  | 
(case lookup_def thy c of  | 
| 19392 | 137  | 
SOME (AxClass info) => info  | 
| 21919 | 138  | 
  | NONE => error ("No such axclass: " ^ quote c));
 | 
| 6379 | 139  | 
|
| 19123 | 140  | 
fun class_intros thy =  | 
| 19392 | 141  | 
let  | 
142  | 
fun add_intro c =  | 
|
| 19511 | 143  | 
      (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
 | 
144  | 
val classes = Sign.all_classes thy;  | 
| 19392 | 145  | 
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
 | 
146  | 
|
| 
 
a67343c6ab2a
sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
 
wenzelm 
parents: 
15853 
diff
changeset
 | 
147  | 
|
| 19460 | 148  | 
fun get_params thy pred =  | 
| 19574 | 149  | 
let val params = #2 (get_axclasses thy);  | 
| 19460 | 150  | 
in fold (fn (x, c) => if pred c then cons x else I) params [] end;  | 
151  | 
||
152  | 
fun all_params_of thy S = get_params thy (fn c => Sign.subsort thy (S, [c]));  | 
|
153  | 
||
| 24964 | 154  | 
fun class_of_param thy = AList.lookup (op =) (#2 (get_axclasses thy));  | 
| 19460 | 155  | 
|
| 21919 | 156  | 
|
| 19511 | 157  | 
(* maintain instances *)  | 
| 19503 | 158  | 
|
| 
30364
 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 
wenzelm 
parents: 
30344 
diff
changeset
 | 
159  | 
fun instance_name (a, c) = Long_Name.base_name c ^ "_" ^ Long_Name.base_name a;  | 
| 25486 | 160  | 
|
| 
25597
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
161  | 
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
 | 
162  | 
val map_instances = AxClassData.map o apsnd o apfst;  | 
| 19528 | 163  | 
|
164  | 
||
| 19574 | 165  | 
fun the_classrel thy (c1, c2) =  | 
166  | 
(case AList.lookup (op =) (#1 (get_instances thy)) (c1, c2) of  | 
|
167  | 
SOME th => Thm.transfer thy th  | 
|
| 24920 | 168  | 
  | NONE => error ("Unproven class relation " ^
 | 
169  | 
Syntax.string_of_classrel (ProofContext.init thy) [c1, c2]));  | 
|
| 19574 | 170  | 
|
171  | 
fun put_classrel arg = map_instances (fn (classrel, arities) =>  | 
|
172  | 
(insert (eq_fst op =) arg classrel, arities));  | 
|
| 19503 | 173  | 
|
174  | 
||
| 19574 | 175  | 
fun the_arity thy a (c, Ss) =  | 
| 
24929
 
408becab067e
renamed AxClass.get_definition to AxClass.get_info (again);
 
wenzelm 
parents: 
24920 
diff
changeset
 | 
176  | 
(case AList.lookup (op =) (Symtab.lookup_list (#2 (get_instances thy)) a) (c, Ss) of  | 
| 33172 | 177  | 
SOME (th, _) => Thm.transfer thy th  | 
| 24920 | 178  | 
  | NONE => error ("Unproven type arity " ^
 | 
179  | 
Syntax.string_of_arity (ProofContext.init thy) (a, Ss, [c])));  | 
|
| 19503 | 180  | 
|
| 33172 | 181  | 
fun thynames_of_arity thy (c, a) =  | 
182  | 
Symtab.lookup_list (#2 (get_instances thy)) a  | 
|
183  | 
|> map_filter (fn ((c', _), (_, name)) => if c = c' then SOME name else NONE)  | 
|
| 27497 | 184  | 
|> rev;  | 
| 27397 | 185  | 
|
| 33172 | 186  | 
fun insert_arity_completions thy (t, ((c, Ss), (th, thy_name))) arities =  | 
| 27547 | 187  | 
let  | 
188  | 
val algebra = Sign.classes_of thy;  | 
|
| 33172 | 189  | 
val super_class_completions =  | 
190  | 
Sign.super_classes thy c  | 
|
| 27547 | 191  | 
|> filter_out (fn c1 => exists (fn ((c2, Ss2), _) => c1 = c2  | 
| 33172 | 192  | 
andalso Sorts.sorts_le algebra (Ss2, Ss)) (Symtab.lookup_list arities t));  | 
| 27554 | 193  | 
val completions = map (fn c1 => (Sorts.weaken algebra  | 
| 27547 | 194  | 
(fn (th, c2) => fn c3 => th RS the_classrel thy (c2, c3)) (th, c) c1  | 
| 33172 | 195  | 
|> Thm.close_derivation, c1)) super_class_completions;  | 
196  | 
val arities' = fold (fn (th1, c1) => Symtab.cons_list (t, ((c1, Ss), (th1, thy_name))))  | 
|
| 27547 | 197  | 
completions arities;  | 
| 33172 | 198  | 
in (null completions, arities') end;  | 
| 19503 | 199  | 
|
| 27547 | 200  | 
fun put_arity ((t, Ss, c), th) thy =  | 
201  | 
let  | 
|
| 33172 | 202  | 
val arity' = (t, ((c, Ss), (th, Context.theory_name thy)));  | 
| 27547 | 203  | 
in  | 
204  | 
thy  | 
|
205  | 
|> map_instances (fn (classrel, arities) => (classrel,  | 
|
206  | 
arities  | 
|
207  | 
|> Symtab.insert_list (eq_fst op =) arity'  | 
|
208  | 
|> insert_arity_completions thy arity'  | 
|
209  | 
|> snd))  | 
|
210  | 
end;  | 
|
| 19503 | 211  | 
|
| 31944 | 212  | 
fun complete_arities thy =  | 
| 27547 | 213  | 
let  | 
214  | 
val arities = snd (get_instances thy);  | 
|
| 33172 | 215  | 
val (finished, arities') = arities  | 
216  | 
|> fold_map (insert_arity_completions thy) (Symtab.dest_list arities);  | 
|
217  | 
in  | 
|
218  | 
if forall I finished then NONE  | 
|
| 27547 | 219  | 
else SOME (thy |> map_instances (fn (classrel, _) => (classrel, arities')))  | 
220  | 
end;  | 
|
221  | 
||
222  | 
val _ = Context.>> (Context.map_theory (Theory.at_begin complete_arities));  | 
|
| 27397 | 223  | 
|
224  | 
||
| 
25597
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
225  | 
(* maintain instance parameters *)  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
226  | 
|
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
227  | 
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
 | 
228  | 
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
 | 
229  | 
|
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
230  | 
fun get_inst_param thy (c, tyco) =  | 
| 
30243
 
09d5944e224e
explicit error message for `improper` instances lacking explicit instance parameter constants
 
haftmann 
parents: 
29579 
diff
changeset
 | 
231  | 
case Symtab.lookup ((the_default Symtab.empty o Symtab.lookup (fst (get_inst_params thy))) c) tyco  | 
| 
 
09d5944e224e
explicit error message for `improper` instances lacking explicit instance parameter constants
 
haftmann 
parents: 
29579 
diff
changeset
 | 
232  | 
of SOME c' => c'  | 
| 
 
09d5944e224e
explicit error message for `improper` instances lacking explicit instance parameter constants
 
haftmann 
parents: 
29579 
diff
changeset
 | 
233  | 
    | NONE => error ("No instance parameter for constant " ^ quote c
 | 
| 
 
09d5944e224e
explicit error message for `improper` instances lacking explicit instance parameter constants
 
haftmann 
parents: 
29579 
diff
changeset
 | 
234  | 
^ " on type constructor " ^ quote tyco);  | 
| 
25597
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
235  | 
|
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
236  | 
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
 | 
237  | 
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
 | 
238  | 
#> (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
 | 
239  | 
|
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
240  | 
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
 | 
241  | 
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
 | 
242  | 
|
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
243  | 
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
 | 
244  | 
(get_inst_params thy) [];  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
245  | 
|
| 31249 | 246  | 
fun get_inst_tyco consts = try (fst o dest_Type o the_single o Consts.typargs consts);  | 
| 
25597
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
247  | 
|
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
248  | 
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
 | 
249  | 
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
 | 
250  | 
|
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
251  | 
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
 | 
252  | 
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
 | 
253  | 
|
| 31249 | 254  | 
fun lookup_inst_param consts params (c, T) = case get_inst_tyco consts (c, T)  | 
255  | 
of SOME tyco => AList.lookup (op =) params (c, tyco)  | 
|
256  | 
| NONE => NONE;  | 
|
257  | 
||
| 
25597
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
258  | 
fun unoverload_const thy (c_ty as (c, _)) =  | 
| 33969 | 259  | 
if is_some (class_of_param thy c)  | 
260  | 
then case get_inst_tyco (Sign.consts_of thy) c_ty  | 
|
261  | 
of SOME tyco => try (param_of_inst thy) (c, tyco) |> the_default c  | 
|
262  | 
| NONE => c  | 
|
263  | 
else c;  | 
|
| 
25597
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
264  | 
|
| 
404
 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 
wenzelm 
parents:  
diff
changeset
 | 
265  | 
|
| 19511 | 266  | 
(** instances **)  | 
| 
19418
 
03b01c9314fc
reworded add_axclass(_i): canonical specification format,
 
wenzelm 
parents: 
19405 
diff
changeset
 | 
267  | 
|
| 
 
03b01c9314fc
reworded add_axclass(_i): canonical specification format,
 
wenzelm 
parents: 
19405 
diff
changeset
 | 
268  | 
(* class relations *)  | 
| 
 
03b01c9314fc
reworded add_axclass(_i): canonical specification format,
 
wenzelm 
parents: 
19405 
diff
changeset
 | 
269  | 
|
| 19405 | 270  | 
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
 | 
271  | 
let  | 
| 
26939
 
1035c89b4c02
moved global pretty/string_of functions from Sign to Syntax;
 
wenzelm 
parents: 
26516 
diff
changeset
 | 
272  | 
val string_of_sort = Syntax.string_of_sort_global thy;  | 
| 19405 | 273  | 
val (c1, c2) = pairself (Sign.certify_class thy) raw_rel;  | 
| 24271 | 274  | 
val _ = Sign.primitive_classrel (c1, c2) (Theory.copy thy);  | 
| 19405 | 275  | 
val _ =  | 
| 19460 | 276  | 
(case subtract (op =) (all_params_of thy [c1]) (all_params_of thy [c2]) of  | 
| 19405 | 277  | 
[] => ()  | 
| 
26939
 
1035c89b4c02
moved global pretty/string_of functions from Sign to Syntax;
 
wenzelm 
parents: 
26516 
diff
changeset
 | 
278  | 
      | xs => raise TYPE ("Class " ^ string_of_sort [c1] ^ " lacks parameter(s) " ^
 | 
| 
 
1035c89b4c02
moved global pretty/string_of functions from Sign to Syntax;
 
wenzelm 
parents: 
26516 
diff
changeset
 | 
279  | 
commas_quote xs ^ " of " ^ string_of_sort [c2], [], []));  | 
| 19405 | 280  | 
in (c1, c2) end;  | 
281  | 
||
282  | 
fun read_classrel thy raw_rel =  | 
|
283  | 
cert_classrel thy (pairself (Sign.read_class thy) raw_rel)  | 
|
284  | 
handle TYPE (msg, _, _) => error msg;  | 
|
285  | 
||
286  | 
||
| 
25597
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
287  | 
(* declaration and definition of instances of overloaded constants *)  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
288  | 
|
| 31249 | 289  | 
fun inst_tyco_of thy (c, T) = case get_inst_tyco (Sign.consts_of thy) (c, T)  | 
290  | 
of SOME tyco => tyco  | 
|
291  | 
  | NONE => error ("Illegal type for instantiation of class parameter: "
 | 
|
292  | 
^ quote (c ^ " :: " ^ Syntax.string_of_typ_global thy T));  | 
|
293  | 
||
| 
25597
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
294  | 
fun declare_overloaded (c, T) thy =  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
295  | 
let  | 
| 25605 | 296  | 
val class = case class_of_param thy c  | 
297  | 
of SOME class => class  | 
|
298  | 
      | NONE => error ("Not a class parameter: " ^ quote c);
 | 
|
| 31249 | 299  | 
val tyco = inst_tyco_of thy (c, T);  | 
| 
25597
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
300  | 
val name_inst = instance_name (tyco, class) ^ "_inst";  | 
| 
30364
 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 
wenzelm 
parents: 
30344 
diff
changeset
 | 
301  | 
val c' = Long_Name.base_name c ^ "_" ^ Long_Name.base_name tyco;  | 
| 
25597
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
302  | 
val T' = Type.strip_sorts T;  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
303  | 
in  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
304  | 
thy  | 
| 30469 | 305  | 
|> Sign.mandatory_path name_inst  | 
| 
33173
 
b8ca12f6681a
eliminated obsolete tags for types/consts -- now handled via name space, in strongly typed fashion;
 
wenzelm 
parents: 
33172 
diff
changeset
 | 
306  | 
|> Sign.declare_const ((Binding.name c', T'), NoSyn)  | 
| 
30344
 
10a67c5ddddb
more uniform handling of binding in targets and derived elements;
 
wenzelm 
parents: 
30280 
diff
changeset
 | 
307  | 
|-> (fn const' as Const (c'', _) =>  | 
| 
 
10a67c5ddddb
more uniform handling of binding in targets and derived elements;
 
wenzelm 
parents: 
30280 
diff
changeset
 | 
308  | 
Thm.add_def false true  | 
| 
 
10a67c5ddddb
more uniform handling of binding in targets and derived elements;
 
wenzelm 
parents: 
30280 
diff
changeset
 | 
309  | 
(Binding.name (Thm.def_name c'), Logic.mk_equals (Const (c, T'), const'))  | 
| 
 
10a67c5ddddb
more uniform handling of binding in targets and derived elements;
 
wenzelm 
parents: 
30280 
diff
changeset
 | 
310  | 
#>> Thm.varifyT  | 
| 
 
10a67c5ddddb
more uniform handling of binding in targets and derived elements;
 
wenzelm 
parents: 
30280 
diff
changeset
 | 
311  | 
#-> (fn thm => add_inst_param (c, tyco) (c'', thm)  | 
| 
33315
 
784c1b09d485
eliminated obsolete/unused Thm.kind_internal/is_internal etc.;
 
wenzelm 
parents: 
33278 
diff
changeset
 | 
312  | 
#> PureThy.add_thm ((Binding.conceal (Binding.name c'), thm), [])  | 
| 
30344
 
10a67c5ddddb
more uniform handling of binding in targets and derived elements;
 
wenzelm 
parents: 
30280 
diff
changeset
 | 
313  | 
#> snd  | 
| 
 
10a67c5ddddb
more uniform handling of binding in targets and derived elements;
 
wenzelm 
parents: 
30280 
diff
changeset
 | 
314  | 
#> Sign.restore_naming thy  | 
| 
 
10a67c5ddddb
more uniform handling of binding in targets and derived elements;
 
wenzelm 
parents: 
30280 
diff
changeset
 | 
315  | 
#> pair (Const (c, T))))  | 
| 
25597
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
316  | 
end;  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
317  | 
|
| 
30519
 
c05c0199826f
coherent binding policy with primitive target operations
 
haftmann 
parents: 
30469 
diff
changeset
 | 
318  | 
fun define_overloaded b (c, t) thy =  | 
| 
25597
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
319  | 
let  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
320  | 
val T = Term.fastype_of t;  | 
| 31249 | 321  | 
val tyco = inst_tyco_of thy (c, T);  | 
| 
25597
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
322  | 
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
 | 
323  | 
val prop = Logic.mk_equals (Const (c', T), t);  | 
| 
30519
 
c05c0199826f
coherent binding policy with primitive target operations
 
haftmann 
parents: 
30469 
diff
changeset
 | 
324  | 
val b' = Thm.def_binding_optional  | 
| 
 
c05c0199826f
coherent binding policy with primitive target operations
 
haftmann 
parents: 
30469 
diff
changeset
 | 
325  | 
(Binding.name (Long_Name.base_name c ^ "_" ^ Long_Name.base_name tyco)) b;  | 
| 
25597
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
326  | 
in  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
327  | 
thy  | 
| 
30519
 
c05c0199826f
coherent binding policy with primitive target operations
 
haftmann 
parents: 
30469 
diff
changeset
 | 
328  | 
|> Thm.add_def false false (b', prop)  | 
| 
25597
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
329  | 
|>> (fn thm => Drule.transitive_thm OF [eq, thm])  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
330  | 
end;  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
331  | 
|
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
332  | 
|
| 
30951
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
333  | 
(* primitive rules *)  | 
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
334  | 
|
| 
31948
 
ea8c8bf47ce3
add_classrel/arity: strip_shyps of stored result;
 
wenzelm 
parents: 
31944 
diff
changeset
 | 
335  | 
fun add_classrel raw_th thy =  | 
| 
30951
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
336  | 
let  | 
| 
31948
 
ea8c8bf47ce3
add_classrel/arity: strip_shyps of stored result;
 
wenzelm 
parents: 
31944 
diff
changeset
 | 
337  | 
val th = Thm.strip_shyps (Thm.transfer thy raw_th);  | 
| 
 
ea8c8bf47ce3
add_classrel/arity: strip_shyps of stored result;
 
wenzelm 
parents: 
31944 
diff
changeset
 | 
338  | 
val prop = Thm.plain_prop_of th;  | 
| 
30951
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
339  | 
    fun err () = raise THM ("add_classrel: malformed class relation", 0, [th]);
 | 
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
340  | 
val rel = Logic.dest_classrel prop handle TERM _ => err ();  | 
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
341  | 
val (c1, c2) = cert_classrel thy rel handle TYPE _ => err ();  | 
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
342  | 
in  | 
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
343  | 
thy  | 
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
344  | 
|> Sign.primitive_classrel (c1, c2)  | 
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
345  | 
|> put_classrel ((c1, c2), Thm.close_derivation (Drule.unconstrainTs th))  | 
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
346  | 
|> perhaps complete_arities  | 
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
347  | 
end;  | 
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
348  | 
|
| 
31948
 
ea8c8bf47ce3
add_classrel/arity: strip_shyps of stored result;
 
wenzelm 
parents: 
31944 
diff
changeset
 | 
349  | 
fun add_arity raw_th thy =  | 
| 
30951
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
350  | 
let  | 
| 
31948
 
ea8c8bf47ce3
add_classrel/arity: strip_shyps of stored result;
 
wenzelm 
parents: 
31944 
diff
changeset
 | 
351  | 
val th = Thm.strip_shyps (Thm.transfer thy raw_th);  | 
| 
 
ea8c8bf47ce3
add_classrel/arity: strip_shyps of stored result;
 
wenzelm 
parents: 
31944 
diff
changeset
 | 
352  | 
val prop = Thm.plain_prop_of th;  | 
| 
30951
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
353  | 
    fun err () = raise THM ("add_arity: malformed type arity", 0, [th]);
 | 
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
354  | 
val (t, Ss, c) = Logic.dest_arity prop handle TERM _ => err ();  | 
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
355  | 
val T = Type (t, map TFree (Name.names Name.context Name.aT Ss));  | 
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
356  | 
val missing_params = Sign.complete_sort thy [c]  | 
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
357  | 
|> maps (these o Option.map #params o try (get_info thy))  | 
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
358  | 
|> filter_out (fn (const, _) => can (get_inst_param thy) (const, t))  | 
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
359  | 
|> (map o apsnd o map_atyps) (K T);  | 
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
360  | 
val _ = map (Sign.certify_sort thy) Ss = Ss orelse err ();  | 
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
361  | 
in  | 
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
362  | 
thy  | 
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
363  | 
|> fold (snd oo declare_overloaded) missing_params  | 
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
364  | 
|> Sign.primitive_arity (t, Ss, [c])  | 
| 33065 | 365  | 
|> put_arity ((t, Ss, c), Thm.close_derivation (Drule.unconstrainTs th))  | 
| 
30951
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
366  | 
end;  | 
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
367  | 
|
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
368  | 
|
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
369  | 
(* tactical proofs *)  | 
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
370  | 
|
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
371  | 
fun prove_classrel raw_rel tac thy =  | 
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
372  | 
let  | 
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
373  | 
val ctxt = ProofContext.init thy;  | 
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
374  | 
val (c1, c2) = cert_classrel thy raw_rel;  | 
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
375  | 
val th = Goal.prove ctxt [] [] (Logic.mk_classrel (c1, c2)) (K tac) handle ERROR msg =>  | 
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
376  | 
      cat_error msg ("The error(s) above occurred while trying to prove class relation " ^
 | 
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
377  | 
quote (Syntax.string_of_classrel ctxt [c1, c2]));  | 
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
378  | 
in  | 
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
379  | 
thy  | 
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
380  | 
|> PureThy.add_thms [((Binding.name  | 
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
381  | 
(prefix classrel_prefix (Logic.name_classrel (c1, c2))), th), [])]  | 
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
382  | 
|-> (fn [th'] => add_classrel th')  | 
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
383  | 
end;  | 
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
384  | 
|
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
385  | 
fun prove_arity raw_arity tac thy =  | 
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
386  | 
let  | 
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
387  | 
val ctxt = ProofContext.init thy;  | 
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
388  | 
val arity = Sign.cert_arity thy raw_arity;  | 
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
389  | 
val names = map (prefix arity_prefix) (Logic.name_arities arity);  | 
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
390  | 
val props = Logic.mk_arities arity;  | 
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
391  | 
val ths = Goal.prove_multi ctxt [] [] props  | 
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
392  | 
(fn _ => Goal.precise_conjunction_tac (length props) 1 THEN tac) handle ERROR msg =>  | 
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
393  | 
        cat_error msg ("The error(s) above occurred while trying to prove type arity " ^
 | 
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
394  | 
quote (Syntax.string_of_arity ctxt arity));  | 
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
395  | 
in  | 
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
396  | 
thy  | 
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
397  | 
|> PureThy.add_thms (map (rpair []) (map Binding.name names ~~ ths))  | 
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
398  | 
|-> fold add_arity  | 
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
399  | 
end;  | 
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
400  | 
|
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
401  | 
|
| 19511 | 402  | 
|
403  | 
(** class definitions **)  | 
|
404  | 
||
| 23421 | 405  | 
fun split_defined n eq =  | 
406  | 
let  | 
|
407  | 
val intro =  | 
|
408  | 
(eq RS Drule.equal_elim_rule2)  | 
|
409  | 
|> Conjunction.curry_balanced n  | 
|
410  | 
|> n = 0 ? Thm.eq_assumption 1;  | 
|
411  | 
val dests =  | 
|
412  | 
if n = 0 then []  | 
|
413  | 
else  | 
|
414  | 
(eq RS Drule.equal_elim_rule1)  | 
|
| 32765 | 415  | 
|> Balanced_Tree.dest (fn th =>  | 
| 23421 | 416  | 
(th RS Conjunction.conjunctionD1, th RS Conjunction.conjunctionD2)) n;  | 
417  | 
in (intro, dests) end;  | 
|
418  | 
||
| 24964 | 419  | 
fun define_class (bclass, raw_super) raw_params raw_specs thy =  | 
| 19511 | 420  | 
let  | 
421  | 
val ctxt = ProofContext.init thy;  | 
|
| 24964 | 422  | 
val pp = Syntax.pp ctxt;  | 
| 19511 | 423  | 
|
424  | 
||
| 24964 | 425  | 
(* class *)  | 
| 19511 | 426  | 
|
| 
30344
 
10a67c5ddddb
more uniform handling of binding in targets and derived elements;
 
wenzelm 
parents: 
30280 
diff
changeset
 | 
427  | 
val bconst = Binding.map_name Logic.const_of_class bclass;  | 
| 
 
10a67c5ddddb
more uniform handling of binding in targets and derived elements;
 
wenzelm 
parents: 
30280 
diff
changeset
 | 
428  | 
val class = Sign.full_name thy bclass;  | 
| 24731 | 429  | 
val super = Sign.minimize_sort thy (Sign.certify_sort thy raw_super);  | 
| 19511 | 430  | 
|
| 24964 | 431  | 
fun check_constraint (a, S) =  | 
432  | 
if Sign.subsort thy (super, S) then ()  | 
|
433  | 
      else error ("Sort constraint of type variable " ^
 | 
|
| 
32966
 
5b21661fe618
indicate CRITICAL nature of various setmp combinators;
 
wenzelm 
parents: 
32791 
diff
changeset
 | 
434  | 
setmp_CRITICAL show_sorts true (Pretty.string_of_typ pp) (TFree (a, S)) ^  | 
| 24964 | 435  | 
" needs to be weaker than " ^ Pretty.string_of_sort pp super);  | 
436  | 
||
437  | 
||
438  | 
(* params *)  | 
|
439  | 
||
440  | 
val params = raw_params |> map (fn p =>  | 
|
441  | 
let  | 
|
442  | 
val T = Sign.the_const_type thy p;  | 
|
443  | 
val _ =  | 
|
444  | 
(case Term.add_tvarsT T [] of  | 
|
445  | 
[((a, _), S)] => check_constraint (a, S)  | 
|
446  | 
          | _ => error ("Exactly one type variable expected in class parameter " ^ quote p));
 | 
|
447  | 
val T' = Term.map_type_tvar (fn _ => TFree (Name.aT, [class])) T;  | 
|
448  | 
in (p, T') end);  | 
|
449  | 
||
450  | 
||
451  | 
(* axioms *)  | 
|
452  | 
||
| 19511 | 453  | 
fun prep_axiom t =  | 
454  | 
(case Term.add_tfrees t [] of  | 
|
| 24964 | 455  | 
[(a, S)] => check_constraint (a, S)  | 
456  | 
| [] => ()  | 
|
457  | 
      | _ => error ("Multiple type variables in class axiom:\n" ^ Pretty.string_of_term pp t);
 | 
|
458  | 
t  | 
|
459  | 
|> Term.map_types (Term.map_atyps (fn TFree _ => Term.aT [] | U => U))  | 
|
460  | 
|> Logic.close_form);  | 
|
| 19511 | 461  | 
|
| 24681 | 462  | 
val axiomss = map (map (prep_axiom o Sign.cert_prop thy) o snd) raw_specs;  | 
| 22745 | 463  | 
val name_atts = map fst raw_specs;  | 
| 19511 | 464  | 
|
465  | 
||
466  | 
(* definition *)  | 
|
467  | 
||
| 
31943
 
5e960a0780a2
renamed inclass/Inclass to of_class/OfClass, in accordance to of_sort;
 
wenzelm 
parents: 
31904 
diff
changeset
 | 
468  | 
val conjs = map (curry Logic.mk_of_class (Term.aT [])) super @ flat axiomss;  | 
| 19511 | 469  | 
val class_eq =  | 
| 
31943
 
5e960a0780a2
renamed inclass/Inclass to of_class/OfClass, in accordance to of_sort;
 
wenzelm 
parents: 
31904 
diff
changeset
 | 
470  | 
Logic.mk_equals (Logic.mk_of_class (Term.aT [], class), Logic.mk_conjunction_balanced conjs);  | 
| 19511 | 471  | 
|
472  | 
val ([def], def_thy) =  | 
|
473  | 
thy  | 
|
474  | 
|> Sign.primitive_class (bclass, super)  | 
|
| 
30344
 
10a67c5ddddb
more uniform handling of binding in targets and derived elements;
 
wenzelm 
parents: 
30280 
diff
changeset
 | 
475  | 
|> PureThy.add_defs false [((Binding.map_name Thm.def_name bconst, class_eq), [])];  | 
| 19511 | 476  | 
val (raw_intro, (raw_classrel, raw_axioms)) =  | 
| 23421 | 477  | 
split_defined (length conjs) def ||> chop (length super);  | 
| 19392 | 478  | 
|
| 
19418
 
03b01c9314fc
reworded add_axclass(_i): canonical specification format,
 
wenzelm 
parents: 
19405 
diff
changeset
 | 
479  | 
|
| 19511 | 480  | 
(* facts *)  | 
481  | 
||
482  | 
val class_triv = Thm.class_triv def_thy class;  | 
|
483  | 
val ([(_, [intro]), (_, classrel), (_, axioms)], facts_thy) =  | 
|
484  | 
def_thy  | 
|
| 30469 | 485  | 
|> Sign.mandatory_path (Binding.name_of bconst)  | 
| 
27691
 
ce171cbd4b93
PureThy: dropped note_thmss_qualified, dropped _i suffix
 
haftmann 
parents: 
27683 
diff
changeset
 | 
486  | 
|> PureThy.note_thmss ""  | 
| 
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: 
34259 
diff
changeset
 | 
487  | 
[((Binding.name introN, []), [([Drule.export_without_context raw_intro], [])]),  | 
| 
 
c839a4c670c6
renamed old-style Drule.standard to Drule.export_without_context, to emphasize that this is in no way a standard operation;
 
wenzelm 
parents: 
34259 
diff
changeset
 | 
488  | 
((Binding.name superN, []), [(map Drule.export_without_context raw_classrel, [])]),  | 
| 28965 | 489  | 
((Binding.name axiomsN, []),  | 
| 
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: 
34259 
diff
changeset
 | 
490  | 
[(map (fn th => Drule.export_without_context (class_triv RS th)) raw_axioms, [])])]  | 
| 
27691
 
ce171cbd4b93
PureThy: dropped note_thmss_qualified, dropped _i suffix
 
haftmann 
parents: 
27683 
diff
changeset
 | 
491  | 
||> Sign.restore_naming def_thy;  | 
| 19511 | 492  | 
|
| 24847 | 493  | 
|
| 19511 | 494  | 
(* result *)  | 
495  | 
||
| 24964 | 496  | 
val axclass = make_axclass ((def, intro, axioms), params);  | 
| 19511 | 497  | 
val result_thy =  | 
498  | 
facts_thy  | 
|
| 19574 | 499  | 
|> fold put_classrel (map (pair class) super ~~ classrel)  | 
| 
30344
 
10a67c5ddddb
more uniform handling of binding in targets and derived elements;
 
wenzelm 
parents: 
30280 
diff
changeset
 | 
500  | 
|> Sign.add_path (Binding.name_of bconst)  | 
| 
27691
 
ce171cbd4b93
PureThy: dropped note_thmss_qualified, dropped _i suffix
 
haftmann 
parents: 
27683 
diff
changeset
 | 
501  | 
|> PureThy.note_thmss "" (name_atts ~~ map Thm.simple_fact (unflat axiomss axioms)) |> snd  | 
| 19511 | 502  | 
|> Sign.restore_naming facts_thy  | 
| 19574 | 503  | 
|> map_axclasses (fn (axclasses, parameters) =>  | 
| 21919 | 504  | 
(Symtab.update (class, axclass) axclasses,  | 
| 24964 | 505  | 
fold (fn (x, _) => add_param pp (x, class)) params parameters));  | 
| 19511 | 506  | 
|
507  | 
in (class, result_thy) end;  | 
|
508  | 
||
509  | 
||
| 19531 | 510  | 
|
| 19511 | 511  | 
(** axiomatizations **)  | 
512  | 
||
513  | 
local  | 
|
514  | 
||
| 20628 | 515  | 
fun axiomatize prep mk name add raw_args thy =  | 
516  | 
let  | 
|
517  | 
val args = prep thy raw_args;  | 
|
518  | 
val specs = mk args;  | 
|
519  | 
val names = name args;  | 
|
| 29579 | 520  | 
in  | 
521  | 
thy  | 
|
522  | 
|> PureThy.add_axioms (map (rpair []) (map Binding.name names ~~ specs))  | 
|
523  | 
|-> fold add  | 
|
524  | 
end;  | 
|
| 19511 | 525  | 
|
526  | 
fun ax_classrel prep =  | 
|
| 20628 | 527  | 
axiomatize (map o prep) (map Logic.mk_classrel)  | 
528  | 
(map (prefix classrel_prefix o Logic.name_classrel)) add_classrel;  | 
|
| 19511 | 529  | 
|
530  | 
fun ax_arity prep =  | 
|
| 20628 | 531  | 
axiomatize prep Logic.mk_arities  | 
532  | 
(map (prefix arity_prefix) o Logic.name_arities) add_arity;  | 
|
| 19511 | 533  | 
|
| 19706 | 534  | 
fun class_const c =  | 
535  | 
(Logic.const_of_class c, Term.itselfT (Term.aT []) --> propT);  | 
|
536  | 
||
| 19511 | 537  | 
fun ax_class prep_class prep_classrel (bclass, raw_super) thy =  | 
538  | 
let  | 
|
| 
30344
 
10a67c5ddddb
more uniform handling of binding in targets and derived elements;
 
wenzelm 
parents: 
30280 
diff
changeset
 | 
539  | 
val class = Sign.full_name thy bclass;  | 
| 24731 | 540  | 
val super = map (prep_class thy) raw_super |> Sign.minimize_sort thy;  | 
| 19511 | 541  | 
in  | 
542  | 
thy  | 
|
543  | 
|> Sign.primitive_class (bclass, super)  | 
|
544  | 
|> ax_classrel prep_classrel (map (fn c => (class, c)) super)  | 
|
| 19706 | 545  | 
|> Theory.add_deps "" (class_const class) (map class_const super)  | 
| 19511 | 546  | 
end;  | 
547  | 
||
548  | 
in  | 
|
549  | 
||
| 
24929
 
408becab067e
renamed AxClass.get_definition to AxClass.get_info (again);
 
wenzelm 
parents: 
24920 
diff
changeset
 | 
550  | 
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
 | 
551  | 
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
 | 
552  | 
val axiomatize_classrel = ax_classrel cert_classrel;  | 
| 
 
408becab067e
renamed AxClass.get_definition to AxClass.get_info (again);
 
wenzelm 
parents: 
24920 
diff
changeset
 | 
553  | 
val axiomatize_classrel_cmd = ax_classrel read_classrel;  | 
| 
 
408becab067e
renamed AxClass.get_definition to AxClass.get_info (again);
 
wenzelm 
parents: 
24920 
diff
changeset
 | 
554  | 
val axiomatize_arity = ax_arity Sign.cert_arity;  | 
| 
 
408becab067e
renamed AxClass.get_definition to AxClass.get_info (again);
 
wenzelm 
parents: 
24920 
diff
changeset
 | 
555  | 
val axiomatize_arity_cmd = ax_arity Sign.read_arity;  | 
| 19511 | 556  | 
|
557  | 
end;  | 
|
558  | 
||
559  | 
||
560  | 
||
561  | 
(** explicit derivations -- cached **)  | 
|
562  | 
||
| 19574 | 563  | 
datatype cache = Types of (class * thm) list Typtab.table;  | 
564  | 
||
| 19511 | 565  | 
local  | 
| 19503 | 566  | 
|
| 19574 | 567  | 
fun lookup_type (Types cache) = AList.lookup (op =) o Typtab.lookup_list cache;  | 
568  | 
fun insert_type T der (Types cache) = Types (Typtab.insert_list (eq_fst op =) (T, der) cache);  | 
|
| 19503 | 569  | 
|
| 19522 | 570  | 
fun derive_type _ (_, []) = []  | 
571  | 
| derive_type thy (typ, sort) =  | 
|
572  | 
let  | 
|
| 31944 | 573  | 
val certT = Thm.ctyp_of thy;  | 
| 19528 | 574  | 
val vars = Term.fold_atyps  | 
| 19522 | 575  | 
(fn T as TFree (_, S) => insert (eq_fst op =) (T, S)  | 
576  | 
| T as TVar (_, S) => insert (eq_fst op =) (T, S)  | 
|
| 19528 | 577  | 
| _ => I) typ [];  | 
| 19574 | 578  | 
val hyps = vars  | 
| 31944 | 579  | 
|> map (fn (T, S) => (T, Thm.of_sort (certT T, S) ~~ S));  | 
| 
26939
 
1035c89b4c02
moved global pretty/string_of functions from Sign to Syntax;
 
wenzelm 
parents: 
26516 
diff
changeset
 | 
580  | 
val ths = (typ, sort)  | 
| 32791 | 581  | 
|> Sorts.of_sort_derivation (Sign.classes_of thy)  | 
| 
22570
 
f166a5416b3f
renamed of_sort_derivation record fields (avoid clash with Alice keywords);
 
wenzelm 
parents: 
22385 
diff
changeset
 | 
582  | 
           {class_relation =
 | 
| 19574 | 583  | 
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
 | 
584  | 
type_constructor =  | 
| 19574 | 585  | 
fn a => fn dom => fn c =>  | 
586  | 
let val Ss = map (map snd) dom and ths = maps (map fst) dom  | 
|
587  | 
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
 | 
588  | 
type_variable =  | 
| 19574 | 589  | 
the_default [] o AList.lookup (op =) hyps};  | 
590  | 
in ths end;  | 
|
| 19503 | 591  | 
|
| 19511 | 592  | 
in  | 
593  | 
||
| 19574 | 594  | 
fun of_sort thy (typ, sort) cache =  | 
595  | 
let  | 
|
596  | 
val sort' = filter (is_none o lookup_type cache typ) sort;  | 
|
597  | 
val ths' = derive_type thy (typ, sort')  | 
|
598  | 
      handle ERROR msg => cat_error msg ("The error(s) above occurred for sort derivation: " ^
 | 
|
| 
26939
 
1035c89b4c02
moved global pretty/string_of functions from Sign to Syntax;
 
wenzelm 
parents: 
26516 
diff
changeset
 | 
599  | 
Syntax.string_of_typ_global thy typ ^ " :: " ^ Syntax.string_of_sort_global thy sort');  | 
| 19574 | 600  | 
val cache' = cache |> fold (insert_type typ) (sort' ~~ ths');  | 
601  | 
val ths =  | 
|
602  | 
sort |> map (fn c =>  | 
|
| 32197 | 603  | 
Goal.finish  | 
604  | 
(Syntax.init_pretty_global thy)  | 
|
605  | 
(the (lookup_type cache' typ c) RS  | 
|
606  | 
Goal.init (Thm.cterm_of thy (Logic.mk_of_class (typ, c))))  | 
|
| 20260 | 607  | 
|> Thm.adjust_maxidx_thm ~1);  | 
| 19574 | 608  | 
in (ths, cache') end;  | 
| 19503 | 609  | 
|
| 19511 | 610  | 
end;  | 
| 
19418
 
03b01c9314fc
reworded add_axclass(_i): canonical specification format,
 
wenzelm 
parents: 
19405 
diff
changeset
 | 
611  | 
|
| 
24929
 
408becab067e
renamed AxClass.get_definition to AxClass.get_info (again);
 
wenzelm 
parents: 
24920 
diff
changeset
 | 
612  | 
val cache = Types Typtab.empty;  | 
| 
 
408becab067e
renamed AxClass.get_definition to AxClass.get_info (again);
 
wenzelm 
parents: 
24920 
diff
changeset
 | 
613  | 
|
| 
15876
 
a67343c6ab2a
sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
 
wenzelm 
parents: 
15853 
diff
changeset
 | 
614  | 
end;  |