| author | wenzelm | 
| Sat, 02 Aug 2014 20:58:15 +0200 | |
| changeset 57845 | a2340800ca1f | 
| parent 56941 | 952833323c99 | 
| child 59058 | a78612c67ec0 | 
| 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  | 
| 36417 | 5  | 
parameters. Proven class relations and type arities.  | 
| 
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  | 
|
| 
51685
 
385ef6706252
more standard module name Axclass (according to file name);
 
wenzelm 
parents: 
51671 
diff
changeset
 | 
8  | 
signature AXCLASS =  | 
| 3938 | 9  | 
sig  | 
| 36427 | 10  | 
  type info = {def: thm, intro: thm, axioms: thm list, params: (string * typ) list}
 | 
11  | 
val get_info: theory -> class -> info  | 
|
12  | 
val class_of_param: theory -> string -> class option  | 
|
13  | 
val instance_name: string * class -> string  | 
|
| 
54931
 
88cf06038e37
uniform orientation of instances as (type constructor, type class)
 
haftmann 
parents: 
54742 
diff
changeset
 | 
14  | 
val thynames_of_arity: theory -> string * class -> string list  | 
| 36427 | 15  | 
val param_of_inst: theory -> string * string -> string  | 
16  | 
val inst_of_param: theory -> string -> (string * string) option  | 
|
17  | 
val unoverload: theory -> thm -> thm  | 
|
18  | 
val overload: theory -> thm -> thm  | 
|
19  | 
val unoverload_conv: theory -> conv  | 
|
20  | 
val overload_conv: theory -> conv  | 
|
21  | 
val lookup_inst_param: Consts.T -> ((string * string) * 'a) list -> string * typ -> 'a option  | 
|
22  | 
val unoverload_const: theory -> string * typ -> string  | 
|
23  | 
val cert_classrel: theory -> class * class -> class * class  | 
|
24  | 
val read_classrel: theory -> xstring * xstring -> class * class  | 
|
25  | 
val declare_overloaded: string * typ -> theory -> term * theory  | 
|
26  | 
val define_overloaded: binding -> string * term -> theory -> thm * theory  | 
|
| 37246 | 27  | 
val add_classrel: thm -> theory -> theory  | 
28  | 
val add_arity: thm -> theory -> theory  | 
|
| 51671 | 29  | 
val prove_classrel: class * class -> (Proof.context -> tactic) -> theory -> theory  | 
30  | 
val prove_arity: string * sort list * sort -> (Proof.context -> tactic) -> theory -> theory  | 
|
| 36427 | 31  | 
val define_class: binding * class list -> string list ->  | 
32  | 
(Thm.binding * term list) list -> theory -> class * theory  | 
|
| 
56941
 
952833323c99
tuned signature to make axiomatizations more easy to spot in the source, via "add_axioms" or "axiomatization";
 
wenzelm 
parents: 
56144 
diff
changeset
 | 
33  | 
val classrel_axiomatization: (class * class) list -> theory -> theory  | 
| 
 
952833323c99
tuned signature to make axiomatizations more easy to spot in the source, via "add_axioms" or "axiomatization";
 
wenzelm 
parents: 
56144 
diff
changeset
 | 
34  | 
val arity_axiomatization: arity -> theory -> theory  | 
| 
 
952833323c99
tuned signature to make axiomatizations more easy to spot in the source, via "add_axioms" or "axiomatization";
 
wenzelm 
parents: 
56144 
diff
changeset
 | 
35  | 
val class_axiomatization: binding * class list -> theory -> theory  | 
| 3938 | 36  | 
end;  | 
| 
404
 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 
wenzelm 
parents:  
diff
changeset
 | 
37  | 
|
| 
51685
 
385ef6706252
more standard module name Axclass (according to file name);
 
wenzelm 
parents: 
51671 
diff
changeset
 | 
38  | 
structure Axclass: AXCLASS =  | 
| 
404
 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 
wenzelm 
parents:  
diff
changeset
 | 
39  | 
struct  | 
| 
 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 
wenzelm 
parents:  
diff
changeset
 | 
40  | 
|
| 19405 | 41  | 
(** theory data **)  | 
| 423 | 42  | 
|
| 
36329
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
43  | 
(* axclass info *)  | 
| 
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
44  | 
|
| 
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
45  | 
type info =  | 
| 
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
46  | 
 {def: thm,
 | 
| 
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
47  | 
intro: thm,  | 
| 
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
48  | 
axioms: thm list,  | 
| 
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
49  | 
params: (string * typ) list};  | 
| 
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
50  | 
|
| 
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
51  | 
fun make_axclass (def, intro, axioms, params): info =  | 
| 
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
52  | 
  {def = def, intro = intro, axioms = axioms, params = params};
 | 
| 
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
53  | 
|
| 
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
54  | 
|
| 19405 | 55  | 
(* class parameters (canonical order) *)  | 
| 423 | 56  | 
|
| 19405 | 57  | 
type param = string * class;  | 
| 423 | 58  | 
|
| 42389 | 59  | 
fun add_param ctxt ((x, c): param) params =  | 
| 19405 | 60  | 
(case AList.lookup (op =) params x of  | 
61  | 
NONE => (x, c) :: params  | 
|
| 
42383
 
0ae4ad40d7b5
simplified pretty printing context, which is only required for certain kernel operations;
 
wenzelm 
parents: 
42375 
diff
changeset
 | 
62  | 
| SOME c' =>  | 
| 42389 | 63  | 
      error ("Duplicate class parameter " ^ quote x ^ " for " ^ Syntax.string_of_sort ctxt [c] ^
 | 
64  | 
(if c = c' then "" else " and " ^ Syntax.string_of_sort ctxt [c'])));  | 
|
| 423 | 65  | 
|
| 
25597
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
66  | 
|
| 19511 | 67  | 
(* setup data *)  | 
| 19392 | 68  | 
|
| 
36329
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
69  | 
datatype data = Data of  | 
| 
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
70  | 
 {axclasses: info Symtab.table,
 | 
| 
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
71  | 
params: param list,  | 
| 
36418
 
752ee03427c2
clarified proven results: store thm only and retrieve proof later via Thm.proof_of (this may also impact parallelism, because internal join_proofs is deferred);
 
wenzelm 
parents: 
36417 
diff
changeset
 | 
72  | 
proven_classrels: thm Symreltab.table,  | 
| 
 
752ee03427c2
clarified proven results: store thm only and retrieve proof later via Thm.proof_of (this may also impact parallelism, because internal join_proofs is deferred);
 
wenzelm 
parents: 
36417 
diff
changeset
 | 
73  | 
proven_arities: ((class * sort list) * (thm * string)) list Symtab.table,  | 
| 
36329
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
74  | 
(*arity theorems with theory name*)  | 
| 
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
75  | 
inst_params:  | 
| 
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
76  | 
(string * thm) Symtab.table Symtab.table *  | 
| 
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
77  | 
(*constant name ~> type constructor ~> (constant name, equation)*)  | 
| 
50764
 
2bbc7ae80634
more precise and complete transitive closure of proven_classrel, using existing Sorts.classes_of which is already closed;
 
wenzelm 
parents: 
50760 
diff
changeset
 | 
78  | 
(string * string) Symtab.table (*constant name ~> (constant name, type constructor)*)};  | 
| 
36329
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
79  | 
|
| 
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
80  | 
fun make_data  | 
| 
50764
 
2bbc7ae80634
more precise and complete transitive closure of proven_classrel, using existing Sorts.classes_of which is already closed;
 
wenzelm 
parents: 
50760 
diff
changeset
 | 
81  | 
(axclasses, params, proven_classrels, proven_arities, inst_params) =  | 
| 
36329
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
82  | 
  Data {axclasses = axclasses, params = params, proven_classrels = proven_classrels,
 | 
| 
50764
 
2bbc7ae80634
more precise and complete transitive closure of proven_classrel, using existing Sorts.classes_of which is already closed;
 
wenzelm 
parents: 
50760 
diff
changeset
 | 
83  | 
proven_arities = proven_arities, inst_params = inst_params};  | 
| 
36329
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
84  | 
|
| 36327 | 85  | 
structure Data = Theory_Data_PP  | 
| 22846 | 86  | 
(  | 
| 
36329
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
87  | 
type T = data;  | 
| 
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
88  | 
val empty =  | 
| 
50764
 
2bbc7ae80634
more precise and complete transitive closure of proven_classrel, using existing Sorts.classes_of which is already closed;
 
wenzelm 
parents: 
50760 
diff
changeset
 | 
89  | 
make_data (Symtab.empty, [], Symreltab.empty, Symtab.empty, (Symtab.empty, Symtab.empty));  | 
| 19574 | 90  | 
val extend = I;  | 
| 
36329
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
91  | 
fun merge pp  | 
| 
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
92  | 
      (Data {axclasses = axclasses1, params = params1, proven_classrels = proven_classrels1,
 | 
| 
50764
 
2bbc7ae80634
more precise and complete transitive closure of proven_classrel, using existing Sorts.classes_of which is already closed;
 
wenzelm 
parents: 
50760 
diff
changeset
 | 
93  | 
proven_arities = proven_arities1, inst_params = inst_params1},  | 
| 
36329
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
94  | 
       Data {axclasses = axclasses2, params = params2, proven_classrels = proven_classrels2,
 | 
| 
50764
 
2bbc7ae80634
more precise and complete transitive closure of proven_classrel, using existing Sorts.classes_of which is already closed;
 
wenzelm 
parents: 
50760 
diff
changeset
 | 
95  | 
proven_arities = proven_arities2, inst_params = inst_params2}) =  | 
| 36325 | 96  | 
let  | 
| 42389 | 97  | 
val ctxt = Syntax.init_pretty pp;  | 
98  | 
||
| 
36329
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
99  | 
val axclasses' = Symtab.merge (K true) (axclasses1, axclasses2);  | 
| 
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
100  | 
val params' =  | 
| 
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
101  | 
if null params1 then params2  | 
| 42389 | 102  | 
else  | 
103  | 
fold_rev (fn p => if member (op =) params1 p then I else add_param ctxt p)  | 
|
104  | 
params2 params1;  | 
|
| 
36329
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
105  | 
|
| 50760 | 106  | 
(*see Theory.at_begin hook for transitive closure of classrels and arity completion*)  | 
| 
50768
 
2172f82de515
tuned -- prefer high-level Table.merge with its slightly more conservative update;
 
wenzelm 
parents: 
50764 
diff
changeset
 | 
107  | 
val proven_classrels' = Symreltab.merge (K true) (proven_classrels1, proven_classrels2);  | 
| 
 
2172f82de515
tuned -- prefer high-level Table.merge with its slightly more conservative update;
 
wenzelm 
parents: 
50764 
diff
changeset
 | 
108  | 
val proven_arities' = Symtab.merge_list (eq_fst op =) (proven_arities1, proven_arities2);  | 
| 
36329
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
109  | 
|
| 
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
110  | 
val inst_params' =  | 
| 
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
111  | 
(Symtab.join (K (Symtab.merge (K true))) (#1 inst_params1, #1 inst_params2),  | 
| 
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
112  | 
Symtab.merge (K true) (#2 inst_params1, #2 inst_params2));  | 
| 36325 | 113  | 
in  | 
| 
50764
 
2bbc7ae80634
more precise and complete transitive closure of proven_classrel, using existing Sorts.classes_of which is already closed;
 
wenzelm 
parents: 
50760 
diff
changeset
 | 
114  | 
make_data (axclasses', params', proven_classrels', proven_arities', inst_params')  | 
| 36325 | 115  | 
end;  | 
| 22846 | 116  | 
);  | 
| 6379 | 117  | 
|
| 
36329
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
118  | 
fun map_data f =  | 
| 
50764
 
2bbc7ae80634
more precise and complete transitive closure of proven_classrel, using existing Sorts.classes_of which is already closed;
 
wenzelm 
parents: 
50760 
diff
changeset
 | 
119  | 
  Data.map (fn Data {axclasses, params, proven_classrels, proven_arities, inst_params} =>
 | 
| 
 
2bbc7ae80634
more precise and complete transitive closure of proven_classrel, using existing Sorts.classes_of which is already closed;
 
wenzelm 
parents: 
50760 
diff
changeset
 | 
120  | 
make_data (f (axclasses, params, proven_classrels, proven_arities, inst_params)));  | 
| 
36329
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
121  | 
|
| 
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
122  | 
fun map_axclasses f =  | 
| 
50764
 
2bbc7ae80634
more precise and complete transitive closure of proven_classrel, using existing Sorts.classes_of which is already closed;
 
wenzelm 
parents: 
50760 
diff
changeset
 | 
123  | 
map_data (fn (axclasses, params, proven_classrels, proven_arities, inst_params) =>  | 
| 
 
2bbc7ae80634
more precise and complete transitive closure of proven_classrel, using existing Sorts.classes_of which is already closed;
 
wenzelm 
parents: 
50760 
diff
changeset
 | 
124  | 
(f axclasses, params, proven_classrels, proven_arities, inst_params));  | 
| 
36329
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
125  | 
|
| 
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
126  | 
fun map_params f =  | 
| 
50764
 
2bbc7ae80634
more precise and complete transitive closure of proven_classrel, using existing Sorts.classes_of which is already closed;
 
wenzelm 
parents: 
50760 
diff
changeset
 | 
127  | 
map_data (fn (axclasses, params, proven_classrels, proven_arities, inst_params) =>  | 
| 
 
2bbc7ae80634
more precise and complete transitive closure of proven_classrel, using existing Sorts.classes_of which is already closed;
 
wenzelm 
parents: 
50760 
diff
changeset
 | 
128  | 
(axclasses, f params, proven_classrels, proven_arities, inst_params));  | 
| 
36329
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
129  | 
|
| 
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
130  | 
fun map_proven_classrels f =  | 
| 
50764
 
2bbc7ae80634
more precise and complete transitive closure of proven_classrel, using existing Sorts.classes_of which is already closed;
 
wenzelm 
parents: 
50760 
diff
changeset
 | 
131  | 
map_data (fn (axclasses, params, proven_classrels, proven_arities, inst_params) =>  | 
| 
 
2bbc7ae80634
more precise and complete transitive closure of proven_classrel, using existing Sorts.classes_of which is already closed;
 
wenzelm 
parents: 
50760 
diff
changeset
 | 
132  | 
(axclasses, params, f proven_classrels, proven_arities, inst_params));  | 
| 
36329
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
133  | 
|
| 
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
134  | 
fun map_proven_arities f =  | 
| 
50764
 
2bbc7ae80634
more precise and complete transitive closure of proven_classrel, using existing Sorts.classes_of which is already closed;
 
wenzelm 
parents: 
50760 
diff
changeset
 | 
135  | 
map_data (fn (axclasses, params, proven_classrels, proven_arities, inst_params) =>  | 
| 
 
2bbc7ae80634
more precise and complete transitive closure of proven_classrel, using existing Sorts.classes_of which is already closed;
 
wenzelm 
parents: 
50760 
diff
changeset
 | 
136  | 
(axclasses, params, proven_classrels, f proven_arities, inst_params));  | 
| 
36329
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
137  | 
|
| 
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
138  | 
fun map_inst_params f =  | 
| 
50764
 
2bbc7ae80634
more precise and complete transitive closure of proven_classrel, using existing Sorts.classes_of which is already closed;
 
wenzelm 
parents: 
50760 
diff
changeset
 | 
139  | 
map_data (fn (axclasses, params, proven_classrels, proven_arities, inst_params) =>  | 
| 
 
2bbc7ae80634
more precise and complete transitive closure of proven_classrel, using existing Sorts.classes_of which is already closed;
 
wenzelm 
parents: 
50760 
diff
changeset
 | 
140  | 
(axclasses, params, proven_classrels, proven_arities, f inst_params));  | 
| 
36329
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
141  | 
|
| 
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
142  | 
val rep_data = Data.get #> (fn Data args => args);  | 
| 
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
143  | 
|
| 
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
144  | 
val axclasses_of = #axclasses o rep_data;  | 
| 
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
145  | 
val params_of = #params o rep_data;  | 
| 
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
146  | 
val proven_classrels_of = #proven_classrels o rep_data;  | 
| 
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
147  | 
val proven_arities_of = #proven_arities o rep_data;  | 
| 
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
148  | 
val inst_params_of = #inst_params o rep_data;  | 
| 
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
149  | 
|
| 6379 | 150  | 
|
| 36346 | 151  | 
(* axclasses with parameters *)  | 
| 19392 | 152  | 
|
| 
24929
 
408becab067e
renamed AxClass.get_definition to AxClass.get_info (again);
 
wenzelm 
parents: 
24920 
diff
changeset
 | 
153  | 
fun get_info thy c =  | 
| 
36329
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
154  | 
(case Symtab.lookup (axclasses_of thy) c of  | 
| 36327 | 155  | 
SOME info => info  | 
| 21919 | 156  | 
  | NONE => error ("No such axclass: " ^ quote c));
 | 
| 6379 | 157  | 
|
| 36327 | 158  | 
fun all_params_of thy S =  | 
| 
36329
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
159  | 
let val params = params_of thy;  | 
| 36327 | 160  | 
in fold (fn (x, c) => if Sign.subsort thy (S, [c]) then cons x else I) params [] end;  | 
| 19460 | 161  | 
|
| 
36329
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
162  | 
fun class_of_param thy = AList.lookup (op =) (params_of thy);  | 
| 19460 | 163  | 
|
| 21919 | 164  | 
|
| 19511 | 165  | 
(* maintain instances *)  | 
| 19503 | 166  | 
|
| 
36329
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
167  | 
val classrel_prefix = "classrel_";  | 
| 
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
168  | 
val arity_prefix = "arity_";  | 
| 25486 | 169  | 
|
| 
36329
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
170  | 
fun instance_name (a, c) = Long_Name.base_name c ^ "_" ^ Long_Name.base_name a;  | 
| 19528 | 171  | 
|
172  | 
||
| 
50764
 
2bbc7ae80634
more precise and complete transitive closure of proven_classrel, using existing Sorts.classes_of which is already closed;
 
wenzelm 
parents: 
50760 
diff
changeset
 | 
173  | 
val update_classrel = map_proven_classrels o Symreltab.update;  | 
| 
36421
 
066e35d1c0d7
tuned classrel completion -- bypass composition with reflexive edges;
 
wenzelm 
parents: 
36420 
diff
changeset
 | 
174  | 
|
| 
50764
 
2bbc7ae80634
more precise and complete transitive closure of proven_classrel, using existing Sorts.classes_of which is already closed;
 
wenzelm 
parents: 
50760 
diff
changeset
 | 
175  | 
val is_classrel = Symreltab.defined o proven_classrels_of;  | 
| 
36421
 
066e35d1c0d7
tuned classrel completion -- bypass composition with reflexive edges;
 
wenzelm 
parents: 
36420 
diff
changeset
 | 
176  | 
|
| 19574 | 177  | 
fun the_classrel thy (c1, c2) =  | 
| 
36329
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
178  | 
(case Symreltab.lookup (proven_classrels_of thy) (c1, c2) of  | 
| 
50764
 
2bbc7ae80634
more precise and complete transitive closure of proven_classrel, using existing Sorts.classes_of which is already closed;
 
wenzelm 
parents: 
50760 
diff
changeset
 | 
179  | 
SOME thm => Thm.transfer thy thm  | 
| 24920 | 180  | 
  | NONE => error ("Unproven class relation " ^
 | 
| 
50764
 
2bbc7ae80634
more precise and complete transitive closure of proven_classrel, using existing Sorts.classes_of which is already closed;
 
wenzelm 
parents: 
50760 
diff
changeset
 | 
181  | 
Syntax.string_of_classrel (Proof_Context.init_global thy) [c1, c2]));  | 
| 36325 | 182  | 
|
183  | 
fun complete_classrels thy =  | 
|
184  | 
let  | 
|
| 
50764
 
2bbc7ae80634
more precise and complete transitive closure of proven_classrel, using existing Sorts.classes_of which is already closed;
 
wenzelm 
parents: 
50760 
diff
changeset
 | 
185  | 
fun complete (c, (_, (all_preds, all_succs))) (finished1, thy1) =  | 
| 
 
2bbc7ae80634
more precise and complete transitive closure of proven_classrel, using existing Sorts.classes_of which is already closed;
 
wenzelm 
parents: 
50760 
diff
changeset
 | 
186  | 
let  | 
| 50769 | 187  | 
fun compl c1 c2 (finished2, thy2) =  | 
| 
50764
 
2bbc7ae80634
more precise and complete transitive closure of proven_classrel, using existing Sorts.classes_of which is already closed;
 
wenzelm 
parents: 
50760 
diff
changeset
 | 
188  | 
if is_classrel thy2 (c1, c2) then (finished2, thy2)  | 
| 
 
2bbc7ae80634
more precise and complete transitive closure of proven_classrel, using existing Sorts.classes_of which is already closed;
 
wenzelm 
parents: 
50760 
diff
changeset
 | 
189  | 
else  | 
| 
 
2bbc7ae80634
more precise and complete transitive closure of proven_classrel, using existing Sorts.classes_of which is already closed;
 
wenzelm 
parents: 
50760 
diff
changeset
 | 
190  | 
(false,  | 
| 
 
2bbc7ae80634
more precise and complete transitive closure of proven_classrel, using existing Sorts.classes_of which is already closed;
 
wenzelm 
parents: 
50760 
diff
changeset
 | 
191  | 
thy2  | 
| 
 
2bbc7ae80634
more precise and complete transitive closure of proven_classrel, using existing Sorts.classes_of which is already closed;
 
wenzelm 
parents: 
50760 
diff
changeset
 | 
192  | 
|> update_classrel ((c1, c2),  | 
| 
 
2bbc7ae80634
more precise and complete transitive closure of proven_classrel, using existing Sorts.classes_of which is already closed;
 
wenzelm 
parents: 
50760 
diff
changeset
 | 
193  | 
(the_classrel thy2 (c1, c) RS the_classrel thy2 (c, c2))  | 
| 
 
2bbc7ae80634
more precise and complete transitive closure of proven_classrel, using existing Sorts.classes_of which is already closed;
 
wenzelm 
parents: 
50760 
diff
changeset
 | 
194  | 
|> Drule.instantiate' [SOME (ctyp_of thy2 (TVar ((Name.aT, 0), [])))] []  | 
| 
 
2bbc7ae80634
more precise and complete transitive closure of proven_classrel, using existing Sorts.classes_of which is already closed;
 
wenzelm 
parents: 
50760 
diff
changeset
 | 
195  | 
|> Thm.close_derivation));  | 
| 
 
2bbc7ae80634
more precise and complete transitive closure of proven_classrel, using existing Sorts.classes_of which is already closed;
 
wenzelm 
parents: 
50760 
diff
changeset
 | 
196  | 
|
| 50769 | 197  | 
val proven = is_classrel thy1;  | 
198  | 
val preds = Graph.Keys.fold (fn c1 => proven (c1, c) ? cons c1) all_preds [];  | 
|
199  | 
val succs = Graph.Keys.fold (fn c2 => proven (c, c2) ? cons c2) all_succs [];  | 
|
200  | 
in  | 
|
201  | 
fold_product compl preds succs (finished1, thy1)  | 
|
202  | 
end;  | 
|
| 36325 | 203  | 
in  | 
| 50769 | 204  | 
(case Graph.fold complete (Sorts.classes_of (Sign.classes_of thy)) (true, thy) of  | 
205  | 
(true, _) => NONE  | 
|
206  | 
| (_, thy') => SOME thy')  | 
|
| 36325 | 207  | 
end;  | 
| 19503 | 208  | 
|
209  | 
||
| 36740 | 210  | 
fun the_arity thy (a, Ss, c) =  | 
| 
36329
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
211  | 
(case AList.lookup (op =) (Symtab.lookup_list (proven_arities_of thy) a) (c, Ss) of  | 
| 
50764
 
2bbc7ae80634
more precise and complete transitive closure of proven_classrel, using existing Sorts.classes_of which is already closed;
 
wenzelm 
parents: 
50760 
diff
changeset
 | 
212  | 
SOME (thm, _) => Thm.transfer thy thm  | 
| 24920 | 213  | 
  | NONE => error ("Unproven type arity " ^
 | 
| 
50764
 
2bbc7ae80634
more precise and complete transitive closure of proven_classrel, using existing Sorts.classes_of which is already closed;
 
wenzelm 
parents: 
50760 
diff
changeset
 | 
214  | 
Syntax.string_of_arity (Proof_Context.init_global thy) (a, Ss, [c])));  | 
| 19503 | 215  | 
|
| 
54931
 
88cf06038e37
uniform orientation of instances as (type constructor, type class)
 
haftmann 
parents: 
54742 
diff
changeset
 | 
216  | 
fun thynames_of_arity thy (a, c) =  | 
| 
36329
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
217  | 
Symtab.lookup_list (proven_arities_of thy) a  | 
| 
36418
 
752ee03427c2
clarified proven results: store thm only and retrieve proof later via Thm.proof_of (this may also impact parallelism, because internal join_proofs is deferred);
 
wenzelm 
parents: 
36417 
diff
changeset
 | 
218  | 
|> map_filter (fn ((c', _), (_, name)) => if c = c' then SOME name else NONE)  | 
| 27497 | 219  | 
|> rev;  | 
| 27397 | 220  | 
|
| 
36419
 
7aea10d10113
tuned aritiy completion -- slightly less intermediate data structures;
 
wenzelm 
parents: 
36418 
diff
changeset
 | 
221  | 
fun insert_arity_completions thy t ((c, Ss), ((th, thy_name))) (finished, arities) =  | 
| 27547 | 222  | 
let  | 
223  | 
val algebra = Sign.classes_of thy;  | 
|
| 36417 | 224  | 
val ars = Symtab.lookup_list arities t;  | 
| 33172 | 225  | 
val super_class_completions =  | 
226  | 
Sign.super_classes thy c  | 
|
| 36417 | 227  | 
|> filter_out (fn c1 => exists (fn ((c2, Ss2), _) =>  | 
228  | 
c1 = c2 andalso Sorts.sorts_le algebra (Ss2, Ss)) ars);  | 
|
229  | 
||
| 
43329
 
84472e198515
tuned signature: Name.invent and Name.invent_names;
 
wenzelm 
parents: 
42389 
diff
changeset
 | 
230  | 
val names = Name.invent Name.context Name.aT (length Ss);  | 
| 36417 | 231  | 
val std_vars = map (fn a => SOME (ctyp_of thy (TVar ((a, 0), [])))) names;  | 
232  | 
||
| 36325 | 233  | 
val completions = super_class_completions |> map (fn c1 =>  | 
234  | 
let  | 
|
| 
36419
 
7aea10d10113
tuned aritiy completion -- slightly less intermediate data structures;
 
wenzelm 
parents: 
36418 
diff
changeset
 | 
235  | 
val th1 =  | 
| 
50764
 
2bbc7ae80634
more precise and complete transitive closure of proven_classrel, using existing Sorts.classes_of which is already closed;
 
wenzelm 
parents: 
50760 
diff
changeset
 | 
236  | 
(th RS the_classrel thy (c, c1))  | 
| 36417 | 237  | 
|> Drule.instantiate' std_vars []  | 
| 36327 | 238  | 
|> Thm.close_derivation;  | 
| 
36418
 
752ee03427c2
clarified proven results: store thm only and retrieve proof later via Thm.proof_of (this may also impact parallelism, because internal join_proofs is deferred);
 
wenzelm 
parents: 
36417 
diff
changeset
 | 
239  | 
in ((th1, thy_name), c1) end);  | 
| 
36419
 
7aea10d10113
tuned aritiy completion -- slightly less intermediate data structures;
 
wenzelm 
parents: 
36418 
diff
changeset
 | 
240  | 
|
| 
 
7aea10d10113
tuned aritiy completion -- slightly less intermediate data structures;
 
wenzelm 
parents: 
36418 
diff
changeset
 | 
241  | 
val finished' = finished andalso null completions;  | 
| 
36418
 
752ee03427c2
clarified proven results: store thm only and retrieve proof later via Thm.proof_of (this may also impact parallelism, because internal join_proofs is deferred);
 
wenzelm 
parents: 
36417 
diff
changeset
 | 
242  | 
val arities' = fold (fn (th, c1) => Symtab.cons_list (t, ((c1, Ss), th))) completions arities;  | 
| 
36419
 
7aea10d10113
tuned aritiy completion -- slightly less intermediate data structures;
 
wenzelm 
parents: 
36418 
diff
changeset
 | 
243  | 
in (finished', arities') end;  | 
| 19503 | 244  | 
|
| 
50764
 
2bbc7ae80634
more precise and complete transitive closure of proven_classrel, using existing Sorts.classes_of which is already closed;
 
wenzelm 
parents: 
50760 
diff
changeset
 | 
245  | 
fun put_arity_completion ((t, Ss, c), th) thy =  | 
| 
36419
 
7aea10d10113
tuned aritiy completion -- slightly less intermediate data structures;
 
wenzelm 
parents: 
36418 
diff
changeset
 | 
246  | 
let val ar = ((c, Ss), (th, Context.theory_name thy)) in  | 
| 27547 | 247  | 
thy  | 
| 
36329
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
248  | 
|> map_proven_arities  | 
| 
36419
 
7aea10d10113
tuned aritiy completion -- slightly less intermediate data structures;
 
wenzelm 
parents: 
36418 
diff
changeset
 | 
249  | 
(Symtab.insert_list (eq_fst op =) (t, ar) #>  | 
| 
 
7aea10d10113
tuned aritiy completion -- slightly less intermediate data structures;
 
wenzelm 
parents: 
36418 
diff
changeset
 | 
250  | 
curry (insert_arity_completions thy t ar) true #> #2)  | 
| 27547 | 251  | 
end;  | 
| 19503 | 252  | 
|
| 31944 | 253  | 
fun complete_arities thy =  | 
| 27547 | 254  | 
let  | 
| 
36329
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
255  | 
val arities = proven_arities_of thy;  | 
| 
36419
 
7aea10d10113
tuned aritiy completion -- slightly less intermediate data structures;
 
wenzelm 
parents: 
36418 
diff
changeset
 | 
256  | 
val (finished, arities') =  | 
| 50769 | 257  | 
Symtab.fold (fn (t, ars) => fold (insert_arity_completions thy t) ars)  | 
258  | 
arities (true, arities);  | 
|
| 33172 | 259  | 
in  | 
| 
36419
 
7aea10d10113
tuned aritiy completion -- slightly less intermediate data structures;
 
wenzelm 
parents: 
36418 
diff
changeset
 | 
260  | 
if finished then NONE  | 
| 
36329
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
261  | 
else SOME (map_proven_arities (K arities') thy)  | 
| 27547 | 262  | 
end;  | 
263  | 
||
| 53171 | 264  | 
val _ = Theory.setup  | 
265  | 
(Theory.at_begin complete_classrels #> Theory.at_begin complete_arities);  | 
|
| 27397 | 266  | 
|
| 36740 | 267  | 
val _ = Proofterm.install_axclass_proofs  | 
268  | 
  {classrel_proof = Thm.proof_of oo the_classrel,
 | 
|
269  | 
arity_proof = Thm.proof_of oo the_arity};  | 
|
| 
36418
 
752ee03427c2
clarified proven results: store thm only and retrieve proof later via Thm.proof_of (this may also impact parallelism, because internal join_proofs is deferred);
 
wenzelm 
parents: 
36417 
diff
changeset
 | 
270  | 
|
| 27397 | 271  | 
|
| 
25597
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
272  | 
(* maintain instance parameters *)  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
273  | 
|
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
274  | 
fun get_inst_param thy (c, tyco) =  | 
| 
36329
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
275  | 
(case Symtab.lookup (the_default Symtab.empty (Symtab.lookup (#1 (inst_params_of thy)) c)) tyco of  | 
| 36327 | 276  | 
SOME c' => c'  | 
277  | 
  | NONE => error ("No instance parameter for constant " ^ quote c ^ " on type " ^ quote tyco));
 | 
|
| 
25597
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
278  | 
|
| 36327 | 279  | 
fun add_inst_param (c, tyco) inst =  | 
280  | 
(map_inst_params o apfst o Symtab.map_default (c, Symtab.empty)) (Symtab.update_new (tyco, inst))  | 
|
| 36417 | 281  | 
#> (map_inst_params o apsnd) (Symtab.update_new (#1 inst, (c, tyco)));  | 
| 
25597
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
282  | 
|
| 
36329
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
283  | 
val inst_of_param = Symtab.lookup o #2 o inst_params_of;  | 
| 36417 | 284  | 
val param_of_inst = #1 oo get_inst_param;  | 
| 
25597
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
285  | 
|
| 36327 | 286  | 
fun inst_thms thy =  | 
| 
36329
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
287  | 
Symtab.fold (Symtab.fold (cons o #2 o #2) o #2) (#1 (inst_params_of thy)) [];  | 
| 
25597
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
288  | 
|
| 36417 | 289  | 
fun get_inst_tyco consts = try (#1 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
 | 
290  | 
|
| 
54742
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
53171 
diff
changeset
 | 
291  | 
fun unoverload thy =  | 
| 
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
53171 
diff
changeset
 | 
292  | 
rewrite_rule (Proof_Context.init_global thy) (inst_thms thy);  | 
| 
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
53171 
diff
changeset
 | 
293  | 
|
| 
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
53171 
diff
changeset
 | 
294  | 
fun overload thy =  | 
| 
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
53171 
diff
changeset
 | 
295  | 
rewrite_rule (Proof_Context.init_global thy) (map Thm.symmetric (inst_thms thy));  | 
| 
25597
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
296  | 
|
| 
54742
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
53171 
diff
changeset
 | 
297  | 
fun unoverload_conv thy =  | 
| 
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
53171 
diff
changeset
 | 
298  | 
Raw_Simplifier.rewrite (Proof_Context.init_global thy) true (inst_thms thy);  | 
| 
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
53171 
diff
changeset
 | 
299  | 
|
| 
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
53171 
diff
changeset
 | 
300  | 
fun overload_conv thy =  | 
| 
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
53171 
diff
changeset
 | 
301  | 
Raw_Simplifier.rewrite (Proof_Context.init_global thy) true (map Thm.symmetric (inst_thms thy));  | 
| 
25597
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
302  | 
|
| 36327 | 303  | 
fun lookup_inst_param consts params (c, T) =  | 
304  | 
(case get_inst_tyco consts (c, T) of  | 
|
305  | 
SOME tyco => AList.lookup (op =) params (c, tyco)  | 
|
306  | 
| NONE => NONE);  | 
|
| 31249 | 307  | 
|
| 
25597
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
308  | 
fun unoverload_const thy (c_ty as (c, _)) =  | 
| 36327 | 309  | 
if is_some (class_of_param thy c) then  | 
310  | 
(case get_inst_tyco (Sign.consts_of thy) c_ty of  | 
|
311  | 
SOME tyco => try (param_of_inst thy) (c, tyco) |> the_default c  | 
|
312  | 
| NONE => c)  | 
|
| 33969 | 313  | 
else c;  | 
| 
25597
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
314  | 
|
| 
404
 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 
wenzelm 
parents:  
diff
changeset
 | 
315  | 
|
| 36327 | 316  | 
|
| 19511 | 317  | 
(** instances **)  | 
| 
19418
 
03b01c9314fc
reworded add_axclass(_i): canonical specification format,
 
wenzelm 
parents: 
19405 
diff
changeset
 | 
318  | 
|
| 
 
03b01c9314fc
reworded add_axclass(_i): canonical specification format,
 
wenzelm 
parents: 
19405 
diff
changeset
 | 
319  | 
(* class relations *)  | 
| 
 
03b01c9314fc
reworded add_axclass(_i): canonical specification format,
 
wenzelm 
parents: 
19405 
diff
changeset
 | 
320  | 
|
| 19405 | 321  | 
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
 | 
322  | 
let  | 
| 
26939
 
1035c89b4c02
moved global pretty/string_of functions from Sign to Syntax;
 
wenzelm 
parents: 
26516 
diff
changeset
 | 
323  | 
val string_of_sort = Syntax.string_of_sort_global thy;  | 
| 19405 | 324  | 
val (c1, c2) = pairself (Sign.certify_class thy) raw_rel;  | 
| 52788 | 325  | 
val _ = Sign.primitive_classrel (c1, c2) thy;  | 
| 19405 | 326  | 
val _ =  | 
| 19460 | 327  | 
(case subtract (op =) (all_params_of thy [c1]) (all_params_of thy [c2]) of  | 
| 19405 | 328  | 
[] => ()  | 
| 
26939
 
1035c89b4c02
moved global pretty/string_of functions from Sign to Syntax;
 
wenzelm 
parents: 
26516 
diff
changeset
 | 
329  | 
      | 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
 | 
330  | 
commas_quote xs ^ " of " ^ string_of_sort [c2], [], []));  | 
| 19405 | 331  | 
in (c1, c2) end;  | 
332  | 
||
333  | 
fun read_classrel thy raw_rel =  | 
|
| 42360 | 334  | 
cert_classrel thy (pairself (Proof_Context.read_class (Proof_Context.init_global thy)) raw_rel)  | 
| 19405 | 335  | 
handle TYPE (msg, _, _) => error msg;  | 
336  | 
||
337  | 
||
| 
25597
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
338  | 
(* declaration and definition of instances of overloaded constants *)  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
339  | 
|
| 
35201
 
c2ddb9395436
axclass: more precise treatment of naming vs. binding;
 
wenzelm 
parents: 
35021 
diff
changeset
 | 
340  | 
fun inst_tyco_of thy (c, T) =  | 
| 
 
c2ddb9395436
axclass: more precise treatment of naming vs. binding;
 
wenzelm 
parents: 
35021 
diff
changeset
 | 
341  | 
(case get_inst_tyco (Sign.consts_of thy) (c, T) of  | 
| 
 
c2ddb9395436
axclass: more precise treatment of naming vs. binding;
 
wenzelm 
parents: 
35021 
diff
changeset
 | 
342  | 
SOME tyco => tyco  | 
| 
 
c2ddb9395436
axclass: more precise treatment of naming vs. binding;
 
wenzelm 
parents: 
35021 
diff
changeset
 | 
343  | 
  | NONE => error ("Illegal type for instantiation of class parameter: " ^
 | 
| 
 
c2ddb9395436
axclass: more precise treatment of naming vs. binding;
 
wenzelm 
parents: 
35021 
diff
changeset
 | 
344  | 
quote (c ^ " :: " ^ Syntax.string_of_typ_global thy T)));  | 
| 31249 | 345  | 
|
| 
25597
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
346  | 
fun declare_overloaded (c, T) thy =  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
347  | 
let  | 
| 
35201
 
c2ddb9395436
axclass: more precise treatment of naming vs. binding;
 
wenzelm 
parents: 
35021 
diff
changeset
 | 
348  | 
val class =  | 
| 
 
c2ddb9395436
axclass: more precise treatment of naming vs. binding;
 
wenzelm 
parents: 
35021 
diff
changeset
 | 
349  | 
(case class_of_param thy c of  | 
| 
 
c2ddb9395436
axclass: more precise treatment of naming vs. binding;
 
wenzelm 
parents: 
35021 
diff
changeset
 | 
350  | 
SOME class => class  | 
| 
 
c2ddb9395436
axclass: more precise treatment of naming vs. binding;
 
wenzelm 
parents: 
35021 
diff
changeset
 | 
351  | 
      | NONE => error ("Not a class parameter: " ^ quote c));
 | 
| 31249 | 352  | 
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
 | 
353  | 
val name_inst = instance_name (tyco, class) ^ "_inst";  | 
| 36417 | 354  | 
val c' = instance_name (tyco, c);  | 
| 
25597
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
355  | 
val T' = Type.strip_sorts T;  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
356  | 
in  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
357  | 
thy  | 
| 
35201
 
c2ddb9395436
axclass: more precise treatment of naming vs. binding;
 
wenzelm 
parents: 
35021 
diff
changeset
 | 
358  | 
|> Sign.qualified_path true (Binding.name name_inst)  | 
| 
42375
 
774df7c59508
report Name_Space.declare/define, relatively to context;
 
wenzelm 
parents: 
42360 
diff
changeset
 | 
359  | 
|> Sign.declare_const_global ((Binding.name c', T'), NoSyn)  | 
| 
30344
 
10a67c5ddddb
more uniform handling of binding in targets and derived elements;
 
wenzelm 
parents: 
30280 
diff
changeset
 | 
360  | 
|-> (fn const' as Const (c'', _) =>  | 
| 
42375
 
774df7c59508
report Name_Space.declare/define, relatively to context;
 
wenzelm 
parents: 
42360 
diff
changeset
 | 
361  | 
Thm.add_def_global false true  | 
| 
30344
 
10a67c5ddddb
more uniform handling of binding in targets and derived elements;
 
wenzelm 
parents: 
30280 
diff
changeset
 | 
362  | 
(Binding.name (Thm.def_name c'), Logic.mk_equals (Const (c, T'), const'))  | 
| 
36106
 
19deea200358
Thm.add_axiom/add_def: return internal name of foundational axiom;
 
wenzelm 
parents: 
35961 
diff
changeset
 | 
363  | 
#>> apsnd Thm.varifyT_global  | 
| 
 
19deea200358
Thm.add_axiom/add_def: return internal name of foundational axiom;
 
wenzelm 
parents: 
35961 
diff
changeset
 | 
364  | 
#-> (fn (_, thm) => add_inst_param (c, tyco) (c'', thm)  | 
| 
39557
 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 
wenzelm 
parents: 
39134 
diff
changeset
 | 
365  | 
#> Global_Theory.add_thm ((Binding.conceal (Binding.name c'), thm), [])  | 
| 36417 | 366  | 
#> #2  | 
| 
36106
 
19deea200358
Thm.add_axiom/add_def: return internal name of foundational axiom;
 
wenzelm 
parents: 
35961 
diff
changeset
 | 
367  | 
#> pair (Const (c, T))))  | 
| 
35201
 
c2ddb9395436
axclass: more precise treatment of naming vs. binding;
 
wenzelm 
parents: 
35021 
diff
changeset
 | 
368  | 
||> Sign.restore_naming thy  | 
| 
25597
 
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  | 
|
| 
30519
 
c05c0199826f
coherent binding policy with primitive target operations
 
haftmann 
parents: 
30469 
diff
changeset
 | 
371  | 
fun define_overloaded b (c, t) thy =  | 
| 
25597
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
372  | 
let  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
373  | 
val T = Term.fastype_of t;  | 
| 31249 | 374  | 
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
 | 
375  | 
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
 | 
376  | 
val prop = Logic.mk_equals (Const (c', T), t);  | 
| 36417 | 377  | 
val b' = Thm.def_binding_optional (Binding.name (instance_name (tyco, c))) b;  | 
| 
25597
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
378  | 
in  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
379  | 
thy  | 
| 
42375
 
774df7c59508
report Name_Space.declare/define, relatively to context;
 
wenzelm 
parents: 
42360 
diff
changeset
 | 
380  | 
|> Thm.add_def_global false false (b', prop)  | 
| 38383 | 381  | 
|>> (fn (_, thm) => Drule.transitive_thm OF [eq, thm])  | 
| 
25597
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
382  | 
end;  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
383  | 
|
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
384  | 
|
| 
30951
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
385  | 
(* primitive rules *)  | 
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
386  | 
|
| 37249 | 387  | 
fun add_classrel raw_th thy =  | 
| 
30951
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
388  | 
let  | 
| 
31948
 
ea8c8bf47ce3
add_classrel/arity: strip_shyps of stored result;
 
wenzelm 
parents: 
31944 
diff
changeset
 | 
389  | 
val th = Thm.strip_shyps (Thm.transfer thy raw_th);  | 
| 
 
ea8c8bf47ce3
add_classrel/arity: strip_shyps of stored result;
 
wenzelm 
parents: 
31944 
diff
changeset
 | 
390  | 
val prop = Thm.plain_prop_of th;  | 
| 
30951
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
391  | 
    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
 | 
392  | 
val rel = Logic.dest_classrel prop handle TERM _ => err ();  | 
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
393  | 
val (c1, c2) = cert_classrel thy rel handle TYPE _ => err ();  | 
| 
56144
 
27167f903c6d
conceal somewhat obscure internal facts, e.g. relevant for 'print_theorems', 'find_theorems';
 
wenzelm 
parents: 
55385 
diff
changeset
 | 
394  | 
val binding =  | 
| 
 
27167f903c6d
conceal somewhat obscure internal facts, e.g. relevant for 'print_theorems', 'find_theorems';
 
wenzelm 
parents: 
55385 
diff
changeset
 | 
395  | 
Binding.conceal (Binding.name (prefix classrel_prefix (Logic.name_classrel (c1, c2))));  | 
| 
 
27167f903c6d
conceal somewhat obscure internal facts, e.g. relevant for 'print_theorems', 'find_theorems';
 
wenzelm 
parents: 
55385 
diff
changeset
 | 
396  | 
val (th', thy') = Global_Theory.store_thm (binding, th) thy;  | 
| 
37232
 
c10fb22a5e0c
classrel and arity theorems are now stored under proper name in theory. add_arity and
 
berghofe 
parents: 
36935 
diff
changeset
 | 
397  | 
val th'' = th'  | 
| 
36934
 
ae0809cff6f0
normalize atyp names after unconstrainT, which may rename atyps arbitrarily;
 
krauss 
parents: 
36881 
diff
changeset
 | 
398  | 
|> Thm.unconstrainT  | 
| 
37232
 
c10fb22a5e0c
classrel and arity theorems are now stored under proper name in theory. add_arity and
 
berghofe 
parents: 
36935 
diff
changeset
 | 
399  | 
|> Drule.instantiate' [SOME (ctyp_of thy' (TVar ((Name.aT, 0), [])))] [];  | 
| 
30951
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
400  | 
in  | 
| 
37232
 
c10fb22a5e0c
classrel and arity theorems are now stored under proper name in theory. add_arity and
 
berghofe 
parents: 
36935 
diff
changeset
 | 
401  | 
thy'  | 
| 
30951
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
402  | 
|> Sign.primitive_classrel (c1, c2)  | 
| 
50764
 
2bbc7ae80634
more precise and complete transitive closure of proven_classrel, using existing Sorts.classes_of which is already closed;
 
wenzelm 
parents: 
50760 
diff
changeset
 | 
403  | 
|> map_proven_classrels (Symreltab.update ((c1, c2), th''))  | 
| 
 
2bbc7ae80634
more precise and complete transitive closure of proven_classrel, using existing Sorts.classes_of which is already closed;
 
wenzelm 
parents: 
50760 
diff
changeset
 | 
404  | 
|> perhaps complete_classrels  | 
| 
30951
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
405  | 
|> perhaps complete_arities  | 
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
406  | 
end;  | 
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
407  | 
|
| 37249 | 408  | 
fun add_arity raw_th thy =  | 
| 
30951
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
409  | 
let  | 
| 
31948
 
ea8c8bf47ce3
add_classrel/arity: strip_shyps of stored result;
 
wenzelm 
parents: 
31944 
diff
changeset
 | 
410  | 
val th = Thm.strip_shyps (Thm.transfer thy raw_th);  | 
| 
 
ea8c8bf47ce3
add_classrel/arity: strip_shyps of stored result;
 
wenzelm 
parents: 
31944 
diff
changeset
 | 
411  | 
val prop = Thm.plain_prop_of th;  | 
| 
30951
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
412  | 
    fun err () = raise THM ("add_arity: malformed type arity", 0, [th]);
 | 
| 
37232
 
c10fb22a5e0c
classrel and arity theorems are now stored under proper name in theory. add_arity and
 
berghofe 
parents: 
36935 
diff
changeset
 | 
413  | 
val arity as (t, Ss, c) = Logic.dest_arity prop handle TERM _ => err ();  | 
| 
 
c10fb22a5e0c
classrel and arity theorems are now stored under proper name in theory. add_arity and
 
berghofe 
parents: 
36935 
diff
changeset
 | 
414  | 
|
| 
56144
 
27167f903c6d
conceal somewhat obscure internal facts, e.g. relevant for 'print_theorems', 'find_theorems';
 
wenzelm 
parents: 
55385 
diff
changeset
 | 
415  | 
val binding =  | 
| 
 
27167f903c6d
conceal somewhat obscure internal facts, e.g. relevant for 'print_theorems', 'find_theorems';
 
wenzelm 
parents: 
55385 
diff
changeset
 | 
416  | 
Binding.conceal (Binding.name (prefix arity_prefix (Logic.name_arity arity)));  | 
| 
 
27167f903c6d
conceal somewhat obscure internal facts, e.g. relevant for 'print_theorems', 'find_theorems';
 
wenzelm 
parents: 
55385 
diff
changeset
 | 
417  | 
val (th', thy') = Global_Theory.store_thm (binding, th) thy;  | 
| 36417 | 418  | 
|
| 
43329
 
84472e198515
tuned signature: Name.invent and Name.invent_names;
 
wenzelm 
parents: 
42389 
diff
changeset
 | 
419  | 
val args = Name.invent_names Name.context Name.aT Ss;  | 
| 36417 | 420  | 
val T = Type (t, map TFree args);  | 
| 
37232
 
c10fb22a5e0c
classrel and arity theorems are now stored under proper name in theory. add_arity and
 
berghofe 
parents: 
36935 
diff
changeset
 | 
421  | 
val std_vars = map (fn (a, S) => SOME (ctyp_of thy' (TVar ((a, 0), [])))) args;  | 
| 36417 | 422  | 
|
| 
37232
 
c10fb22a5e0c
classrel and arity theorems are now stored under proper name in theory. add_arity and
 
berghofe 
parents: 
36935 
diff
changeset
 | 
423  | 
val missing_params = Sign.complete_sort thy' [c]  | 
| 
 
c10fb22a5e0c
classrel and arity theorems are now stored under proper name in theory. add_arity and
 
berghofe 
parents: 
36935 
diff
changeset
 | 
424  | 
|> maps (these o Option.map #params o try (get_info thy'))  | 
| 
 
c10fb22a5e0c
classrel and arity theorems are now stored under proper name in theory. add_arity and
 
berghofe 
parents: 
36935 
diff
changeset
 | 
425  | 
|> filter_out (fn (const, _) => can (get_inst_param thy') (const, t))  | 
| 
30951
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
426  | 
|> (map o apsnd o map_atyps) (K T);  | 
| 
37232
 
c10fb22a5e0c
classrel and arity theorems are now stored under proper name in theory. add_arity and
 
berghofe 
parents: 
36935 
diff
changeset
 | 
427  | 
val th'' = th'  | 
| 
36934
 
ae0809cff6f0
normalize atyp names after unconstrainT, which may rename atyps arbitrarily;
 
krauss 
parents: 
36881 
diff
changeset
 | 
428  | 
|> Thm.unconstrainT  | 
| 
 
ae0809cff6f0
normalize atyp names after unconstrainT, which may rename atyps arbitrarily;
 
krauss 
parents: 
36881 
diff
changeset
 | 
429  | 
|> Drule.instantiate' std_vars [];  | 
| 
30951
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
430  | 
in  | 
| 
37232
 
c10fb22a5e0c
classrel and arity theorems are now stored under proper name in theory. add_arity and
 
berghofe 
parents: 
36935 
diff
changeset
 | 
431  | 
thy'  | 
| 36417 | 432  | 
|> fold (#2 oo declare_overloaded) missing_params  | 
| 
30951
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
433  | 
|> Sign.primitive_arity (t, Ss, [c])  | 
| 
50764
 
2bbc7ae80634
more precise and complete transitive closure of proven_classrel, using existing Sorts.classes_of which is already closed;
 
wenzelm 
parents: 
50760 
diff
changeset
 | 
434  | 
|> put_arity_completion ((t, Ss, c), th'')  | 
| 
30951
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
435  | 
end;  | 
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
436  | 
|
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
437  | 
|
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
438  | 
(* tactical proofs *)  | 
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
439  | 
|
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
440  | 
fun prove_classrel raw_rel tac thy =  | 
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
441  | 
let  | 
| 42360 | 442  | 
val ctxt = Proof_Context.init_global thy;  | 
| 
30951
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
443  | 
val (c1, c2) = cert_classrel thy raw_rel;  | 
| 51671 | 444  | 
val th =  | 
445  | 
      Goal.prove ctxt [] [] (Logic.mk_classrel (c1, c2)) (fn {context, ...} => tac context)
 | 
|
446  | 
handle ERROR msg =>  | 
|
447  | 
          cat_error msg ("The error(s) above occurred while trying to prove class relation " ^
 | 
|
448  | 
quote (Syntax.string_of_classrel ctxt [c1, c2]));  | 
|
| 
30951
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
449  | 
in  | 
| 37249 | 450  | 
thy |> add_classrel th  | 
| 
30951
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
451  | 
end;  | 
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
452  | 
|
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
453  | 
fun prove_arity raw_arity tac thy =  | 
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
454  | 
let  | 
| 42360 | 455  | 
val ctxt = Proof_Context.init_global thy;  | 
456  | 
val arity = Proof_Context.cert_arity ctxt raw_arity;  | 
|
| 
30951
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
457  | 
val names = map (prefix arity_prefix) (Logic.name_arities arity);  | 
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
458  | 
val props = Logic.mk_arities arity;  | 
| 51671 | 459  | 
val ths =  | 
460  | 
Goal.prove_multi ctxt [] [] props  | 
|
461  | 
      (fn {context, ...} => Goal.precise_conjunction_tac (length props) 1 THEN tac context)
 | 
|
462  | 
handle ERROR msg =>  | 
|
463  | 
          cat_error msg ("The error(s) above occurred while trying to prove type arity " ^
 | 
|
464  | 
quote (Syntax.string_of_arity ctxt arity));  | 
|
| 
30951
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
465  | 
in  | 
| 37249 | 466  | 
thy |> fold add_arity ths  | 
| 
30951
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
467  | 
end;  | 
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
468  | 
|
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
469  | 
|
| 19511 | 470  | 
|
471  | 
(** class definitions **)  | 
|
472  | 
||
| 23421 | 473  | 
fun split_defined n eq =  | 
474  | 
let  | 
|
475  | 
val intro =  | 
|
476  | 
(eq RS Drule.equal_elim_rule2)  | 
|
477  | 
|> Conjunction.curry_balanced n  | 
|
478  | 
|> n = 0 ? Thm.eq_assumption 1;  | 
|
479  | 
val dests =  | 
|
480  | 
if n = 0 then []  | 
|
481  | 
else  | 
|
482  | 
(eq RS Drule.equal_elim_rule1)  | 
|
| 32765 | 483  | 
|> Balanced_Tree.dest (fn th =>  | 
| 23421 | 484  | 
(th RS Conjunction.conjunctionD1, th RS Conjunction.conjunctionD2)) n;  | 
485  | 
in (intro, dests) end;  | 
|
486  | 
||
| 24964 | 487  | 
fun define_class (bclass, raw_super) raw_params raw_specs thy =  | 
| 19511 | 488  | 
let  | 
| 
39133
 
70d3915c92f0
pretty printing: prefer regular Proof.context over Pretty.pp, which is mostly for special bootstrap purposes involving theory merge, for example;
 
wenzelm 
parents: 
38383 
diff
changeset
 | 
489  | 
val ctxt = Syntax.init_pretty_global thy;  | 
| 19511 | 490  | 
|
491  | 
||
| 24964 | 492  | 
(* class *)  | 
| 19511 | 493  | 
|
| 
30344
 
10a67c5ddddb
more uniform handling of binding in targets and derived elements;
 
wenzelm 
parents: 
30280 
diff
changeset
 | 
494  | 
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
 | 
495  | 
val class = Sign.full_name thy bclass;  | 
| 24731 | 496  | 
val super = Sign.minimize_sort thy (Sign.certify_sort thy raw_super);  | 
| 19511 | 497  | 
|
| 24964 | 498  | 
fun check_constraint (a, S) =  | 
499  | 
if Sign.subsort thy (super, S) then ()  | 
|
500  | 
      else error ("Sort constraint of type variable " ^
 | 
|
| 
39134
 
917b4b6ba3d2
turned show_sorts/show_types into proper configuration options;
 
wenzelm 
parents: 
39133 
diff
changeset
 | 
501  | 
Syntax.string_of_typ (Config.put show_sorts true ctxt) (TFree (a, S)) ^  | 
| 
39133
 
70d3915c92f0
pretty printing: prefer regular Proof.context over Pretty.pp, which is mostly for special bootstrap purposes involving theory merge, for example;
 
wenzelm 
parents: 
38383 
diff
changeset
 | 
502  | 
" needs to be weaker than " ^ Syntax.string_of_sort ctxt super);  | 
| 24964 | 503  | 
|
504  | 
||
505  | 
(* params *)  | 
|
506  | 
||
507  | 
val params = raw_params |> map (fn p =>  | 
|
508  | 
let  | 
|
509  | 
val T = Sign.the_const_type thy p;  | 
|
510  | 
val _ =  | 
|
511  | 
(case Term.add_tvarsT T [] of  | 
|
512  | 
[((a, _), S)] => check_constraint (a, S)  | 
|
513  | 
          | _ => error ("Exactly one type variable expected in class parameter " ^ quote p));
 | 
|
514  | 
val T' = Term.map_type_tvar (fn _ => TFree (Name.aT, [class])) T;  | 
|
515  | 
in (p, T') end);  | 
|
516  | 
||
517  | 
||
518  | 
(* axioms *)  | 
|
519  | 
||
| 19511 | 520  | 
fun prep_axiom t =  | 
521  | 
(case Term.add_tfrees t [] of  | 
|
| 24964 | 522  | 
[(a, S)] => check_constraint (a, S)  | 
523  | 
| [] => ()  | 
|
| 
39133
 
70d3915c92f0
pretty printing: prefer regular Proof.context over Pretty.pp, which is mostly for special bootstrap purposes involving theory merge, for example;
 
wenzelm 
parents: 
38383 
diff
changeset
 | 
524  | 
      | _ => error ("Multiple type variables in class axiom:\n" ^ Syntax.string_of_term ctxt t);
 | 
| 24964 | 525  | 
t  | 
526  | 
|> Term.map_types (Term.map_atyps (fn TFree _ => Term.aT [] | U => U))  | 
|
527  | 
|> Logic.close_form);  | 
|
| 19511 | 528  | 
|
| 24681 | 529  | 
val axiomss = map (map (prep_axiom o Sign.cert_prop thy) o snd) raw_specs;  | 
| 22745 | 530  | 
val name_atts = map fst raw_specs;  | 
| 19511 | 531  | 
|
532  | 
||
533  | 
(* definition *)  | 
|
534  | 
||
| 35854 | 535  | 
val conjs = Logic.mk_of_sort (Term.aT [], super) @ flat axiomss;  | 
| 19511 | 536  | 
val class_eq =  | 
| 
31943
 
5e960a0780a2
renamed inclass/Inclass to of_class/OfClass, in accordance to of_sort;
 
wenzelm 
parents: 
31904 
diff
changeset
 | 
537  | 
Logic.mk_equals (Logic.mk_of_class (Term.aT [], class), Logic.mk_conjunction_balanced conjs);  | 
| 19511 | 538  | 
|
539  | 
val ([def], def_thy) =  | 
|
540  | 
thy  | 
|
541  | 
|> Sign.primitive_class (bclass, super)  | 
|
| 
39557
 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 
wenzelm 
parents: 
39134 
diff
changeset
 | 
542  | 
|> Global_Theory.add_defs false [((Thm.def_binding bconst, class_eq), [])];  | 
| 19511 | 543  | 
val (raw_intro, (raw_classrel, raw_axioms)) =  | 
| 23421 | 544  | 
split_defined (length conjs) def ||> chop (length super);  | 
| 19392 | 545  | 
|
| 
19418
 
03b01c9314fc
reworded add_axclass(_i): canonical specification format,
 
wenzelm 
parents: 
19405 
diff
changeset
 | 
546  | 
|
| 19511 | 547  | 
(* facts *)  | 
548  | 
||
549  | 
val class_triv = Thm.class_triv def_thy class;  | 
|
550  | 
val ([(_, [intro]), (_, classrel), (_, axioms)], facts_thy) =  | 
|
551  | 
def_thy  | 
|
| 
35201
 
c2ddb9395436
axclass: more precise treatment of naming vs. binding;
 
wenzelm 
parents: 
35021 
diff
changeset
 | 
552  | 
|> Sign.qualified_path true bconst  | 
| 
39557
 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 
wenzelm 
parents: 
39134 
diff
changeset
 | 
553  | 
|> Global_Theory.note_thmss ""  | 
| 36326 | 554  | 
[((Binding.name "intro", []), [([Drule.export_without_context raw_intro], [])]),  | 
555  | 
((Binding.name "super", []), [(map Drule.export_without_context raw_classrel, [])]),  | 
|
556  | 
((Binding.name "axioms", []),  | 
|
| 
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
 | 
557  | 
[(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
 | 
558  | 
||> Sign.restore_naming def_thy;  | 
| 19511 | 559  | 
|
| 24847 | 560  | 
|
| 19511 | 561  | 
(* result *)  | 
562  | 
||
| 
36329
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
563  | 
val axclass = make_axclass (def, intro, axioms, params);  | 
| 19511 | 564  | 
val result_thy =  | 
565  | 
facts_thy  | 
|
| 
50764
 
2bbc7ae80634
more precise and complete transitive closure of proven_classrel, using existing Sorts.classes_of which is already closed;
 
wenzelm 
parents: 
50760 
diff
changeset
 | 
566  | 
|> map_proven_classrels  | 
| 
 
2bbc7ae80634
more precise and complete transitive closure of proven_classrel, using existing Sorts.classes_of which is already closed;
 
wenzelm 
parents: 
50760 
diff
changeset
 | 
567  | 
(fold2 (fn c => fn th => Symreltab.update ((class, c), th)) super classrel)  | 
| 
 
2bbc7ae80634
more precise and complete transitive closure of proven_classrel, using existing Sorts.classes_of which is already closed;
 
wenzelm 
parents: 
50760 
diff
changeset
 | 
568  | 
|> perhaps complete_classrels  | 
| 
35201
 
c2ddb9395436
axclass: more precise treatment of naming vs. binding;
 
wenzelm 
parents: 
35021 
diff
changeset
 | 
569  | 
|> Sign.qualified_path false bconst  | 
| 
39557
 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 
wenzelm 
parents: 
39134 
diff
changeset
 | 
570  | 
|> Global_Theory.note_thmss "" (name_atts ~~ map Thm.simple_fact (unflat axiomss axioms))  | 
| 
 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 
wenzelm 
parents: 
39134 
diff
changeset
 | 
571  | 
|> #2  | 
| 19511 | 572  | 
|> Sign.restore_naming facts_thy  | 
| 
36329
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
573  | 
|> map_axclasses (Symtab.update (class, axclass))  | 
| 42389 | 574  | 
|> map_params (fold (fn (x, _) => add_param ctxt (x, class)) params);  | 
| 19511 | 575  | 
|
576  | 
in (class, result_thy) end;  | 
|
577  | 
||
578  | 
||
| 19531 | 579  | 
|
| 19511 | 580  | 
(** axiomatizations **)  | 
581  | 
||
582  | 
local  | 
|
583  | 
||
| 36417 | 584  | 
(*old-style axioms*)  | 
| 
56941
 
952833323c99
tuned signature to make axiomatizations more easy to spot in the source, via "add_axioms" or "axiomatization";
 
wenzelm 
parents: 
56144 
diff
changeset
 | 
585  | 
fun add_axioms prep mk name add raw_args thy =  | 
| 20628 | 586  | 
let  | 
587  | 
val args = prep thy raw_args;  | 
|
588  | 
val specs = mk args;  | 
|
589  | 
val names = name args;  | 
|
| 29579 | 590  | 
in  | 
591  | 
thy  | 
|
| 
42375
 
774df7c59508
report Name_Space.declare/define, relatively to context;
 
wenzelm 
parents: 
42360 
diff
changeset
 | 
592  | 
|> fold_map Thm.add_axiom_global (map Binding.name names ~~ specs)  | 
| 37249 | 593  | 
|-> fold (add o Drule.export_without_context o snd)  | 
| 29579 | 594  | 
end;  | 
| 19511 | 595  | 
|
| 19706 | 596  | 
fun class_const c =  | 
597  | 
(Logic.const_of_class c, Term.itselfT (Term.aT []) --> propT);  | 
|
598  | 
||
| 
55385
 
169e12bbf9a3
discontinued axiomatic 'classes', 'classrel', 'arities';
 
wenzelm 
parents: 
54931 
diff
changeset
 | 
599  | 
in  | 
| 
 
169e12bbf9a3
discontinued axiomatic 'classes', 'classrel', 'arities';
 
wenzelm 
parents: 
54931 
diff
changeset
 | 
600  | 
|
| 
56941
 
952833323c99
tuned signature to make axiomatizations more easy to spot in the source, via "add_axioms" or "axiomatization";
 
wenzelm 
parents: 
56144 
diff
changeset
 | 
601  | 
val classrel_axiomatization =  | 
| 
 
952833323c99
tuned signature to make axiomatizations more easy to spot in the source, via "add_axioms" or "axiomatization";
 
wenzelm 
parents: 
56144 
diff
changeset
 | 
602  | 
add_axioms (map o cert_classrel) (map Logic.mk_classrel)  | 
| 
55385
 
169e12bbf9a3
discontinued axiomatic 'classes', 'classrel', 'arities';
 
wenzelm 
parents: 
54931 
diff
changeset
 | 
603  | 
(map (prefix classrel_prefix o Logic.name_classrel)) add_classrel;  | 
| 
 
169e12bbf9a3
discontinued axiomatic 'classes', 'classrel', 'arities';
 
wenzelm 
parents: 
54931 
diff
changeset
 | 
604  | 
|
| 
56941
 
952833323c99
tuned signature to make axiomatizations more easy to spot in the source, via "add_axioms" or "axiomatization";
 
wenzelm 
parents: 
56144 
diff
changeset
 | 
605  | 
val arity_axiomatization =  | 
| 
 
952833323c99
tuned signature to make axiomatizations more easy to spot in the source, via "add_axioms" or "axiomatization";
 
wenzelm 
parents: 
56144 
diff
changeset
 | 
606  | 
add_axioms (Proof_Context.cert_arity o Proof_Context.init_global) Logic.mk_arities  | 
| 
55385
 
169e12bbf9a3
discontinued axiomatic 'classes', 'classrel', 'arities';
 
wenzelm 
parents: 
54931 
diff
changeset
 | 
607  | 
(map (prefix arity_prefix) o Logic.name_arities) add_arity;  | 
| 
 
169e12bbf9a3
discontinued axiomatic 'classes', 'classrel', 'arities';
 
wenzelm 
parents: 
54931 
diff
changeset
 | 
608  | 
|
| 
56941
 
952833323c99
tuned signature to make axiomatizations more easy to spot in the source, via "add_axioms" or "axiomatization";
 
wenzelm 
parents: 
56144 
diff
changeset
 | 
609  | 
fun class_axiomatization (bclass, raw_super) thy =  | 
| 19511 | 610  | 
let  | 
| 
30344
 
10a67c5ddddb
more uniform handling of binding in targets and derived elements;
 
wenzelm 
parents: 
30280 
diff
changeset
 | 
611  | 
val class = Sign.full_name thy bclass;  | 
| 
55385
 
169e12bbf9a3
discontinued axiomatic 'classes', 'classrel', 'arities';
 
wenzelm 
parents: 
54931 
diff
changeset
 | 
612  | 
val super = map (Sign.certify_class thy) raw_super |> Sign.minimize_sort thy;  | 
| 19511 | 613  | 
in  | 
614  | 
thy  | 
|
615  | 
|> Sign.primitive_class (bclass, super)  | 
|
| 
56941
 
952833323c99
tuned signature to make axiomatizations more easy to spot in the source, via "add_axioms" or "axiomatization";
 
wenzelm 
parents: 
56144 
diff
changeset
 | 
616  | 
|> classrel_axiomatization (map (fn c => (class, c)) super)  | 
| 
42375
 
774df7c59508
report Name_Space.declare/define, relatively to context;
 
wenzelm 
parents: 
42360 
diff
changeset
 | 
617  | 
|> Theory.add_deps_global "" (class_const class) (map class_const super)  | 
| 19511 | 618  | 
end;  | 
619  | 
||
620  | 
end;  | 
|
621  | 
||
622  | 
end;  |