| author | wenzelm | 
| Mon, 18 Apr 2011 11:13:29 +0200 | |
| changeset 42383 | 0ae4ad40d7b5 | 
| parent 42375 | 774df7c59508 | 
| child 42389 | b2c6033fc7e4 | 
| 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  | 
|
| 
 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 
wenzelm 
parents:  
diff
changeset
 | 
8  | 
signature AX_CLASS =  | 
| 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  | 
|
14  | 
val thynames_of_arity: theory -> class * string -> string list  | 
|
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  | 
|
| 24589 | 29  | 
val prove_classrel: class * class -> tactic -> theory -> theory  | 
30  | 
val prove_arity: string * sort list * sort -> tactic -> theory -> theory  | 
|
| 36427 | 31  | 
val define_class: binding * class list -> string list ->  | 
32  | 
(Thm.binding * term list) list -> theory -> class * theory  | 
|
| 
30344
 
10a67c5ddddb
more uniform handling of binding in targets and derived elements;
 
wenzelm 
parents: 
30280 
diff
changeset
 | 
33  | 
val axiomatize_class: binding * class list -> theory -> theory  | 
| 
 
10a67c5ddddb
more uniform handling of binding in targets and derived elements;
 
wenzelm 
parents: 
30280 
diff
changeset
 | 
34  | 
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
 | 
35  | 
val axiomatize_classrel: (class * class) list -> theory -> theory  | 
| 
 
408becab067e
renamed AxClass.get_definition to AxClass.get_info (again);
 
wenzelm 
parents: 
24920 
diff
changeset
 | 
36  | 
val axiomatize_classrel_cmd: (xstring * xstring) list -> theory -> theory  | 
| 
 
408becab067e
renamed AxClass.get_definition to AxClass.get_info (again);
 
wenzelm 
parents: 
24920 
diff
changeset
 | 
37  | 
val axiomatize_arity: arity -> theory -> theory  | 
| 
 
408becab067e
renamed AxClass.get_definition to AxClass.get_info (again);
 
wenzelm 
parents: 
24920 
diff
changeset
 | 
38  | 
val axiomatize_arity_cmd: xstring * string list * string -> theory -> theory  | 
| 3938 | 39  | 
end;  | 
| 
404
 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 
wenzelm 
parents:  
diff
changeset
 | 
40  | 
|
| 15801 | 41  | 
structure AxClass: AX_CLASS =  | 
| 
404
 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 
wenzelm 
parents:  
diff
changeset
 | 
42  | 
struct  | 
| 
 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 
wenzelm 
parents:  
diff
changeset
 | 
43  | 
|
| 19405 | 44  | 
(** theory data **)  | 
| 423 | 45  | 
|
| 
36329
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
46  | 
(* axclass info *)  | 
| 
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
47  | 
|
| 
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
48  | 
type info =  | 
| 
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
49  | 
 {def: thm,
 | 
| 
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
50  | 
intro: thm,  | 
| 
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
51  | 
axioms: thm list,  | 
| 
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
52  | 
params: (string * typ) list};  | 
| 
 
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  | 
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
 | 
55  | 
  {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
 | 
56  | 
|
| 
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
57  | 
|
| 19405 | 58  | 
(* class parameters (canonical order) *)  | 
| 423 | 59  | 
|
| 19405 | 60  | 
type param = string * class;  | 
| 423 | 61  | 
|
| 19405 | 62  | 
fun add_param pp ((x, c): param) params =  | 
63  | 
(case AList.lookup (op =) params x of  | 
|
64  | 
NONE => (x, c) :: params  | 
|
| 
42383
 
0ae4ad40d7b5
simplified pretty printing context, which is only required for certain kernel operations;
 
wenzelm 
parents: 
42375 
diff
changeset
 | 
65  | 
| SOME c' =>  | 
| 
 
0ae4ad40d7b5
simplified pretty printing context, which is only required for certain kernel operations;
 
wenzelm 
parents: 
42375 
diff
changeset
 | 
66  | 
let val ctxt = Syntax.init_pretty pp in  | 
| 
 
0ae4ad40d7b5
simplified pretty printing context, which is only required for certain kernel operations;
 
wenzelm 
parents: 
42375 
diff
changeset
 | 
67  | 
        error ("Duplicate class parameter " ^ quote x ^
 | 
| 
 
0ae4ad40d7b5
simplified pretty printing context, which is only required for certain kernel operations;
 
wenzelm 
parents: 
42375 
diff
changeset
 | 
68  | 
" for " ^ Syntax.string_of_sort ctxt [c] ^  | 
| 
 
0ae4ad40d7b5
simplified pretty printing context, which is only required for certain kernel operations;
 
wenzelm 
parents: 
42375 
diff
changeset
 | 
69  | 
(if c = c' then "" else " and " ^ Syntax.string_of_sort ctxt [c']))  | 
| 
 
0ae4ad40d7b5
simplified pretty printing context, which is only required for certain kernel operations;
 
wenzelm 
parents: 
42375 
diff
changeset
 | 
70  | 
end);  | 
| 423 | 71  | 
|
| 
25597
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
72  | 
|
| 19511 | 73  | 
(* setup data *)  | 
| 19392 | 74  | 
|
| 
36329
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
75  | 
datatype data = Data of  | 
| 
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
76  | 
 {axclasses: info Symtab.table,
 | 
| 
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
77  | 
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
 | 
78  | 
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
 | 
79  | 
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
 | 
80  | 
(*arity theorems with theory name*)  | 
| 
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
81  | 
inst_params:  | 
| 
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
82  | 
(string * thm) Symtab.table Symtab.table *  | 
| 
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
83  | 
(*constant name ~> type constructor ~> (constant name, equation)*)  | 
| 
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
84  | 
(string * string) Symtab.table (*constant name ~> (constant name, type constructor)*),  | 
| 
36420
 
66fd989c8e15
tuned diff_classrels -- avoid slightly inefficient Symreltab.keys;
 
wenzelm 
parents: 
36419 
diff
changeset
 | 
85  | 
diff_classrels: (class * class) list};  | 
| 
36329
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
86  | 
|
| 
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
87  | 
fun make_data  | 
| 
36420
 
66fd989c8e15
tuned diff_classrels -- avoid slightly inefficient Symreltab.keys;
 
wenzelm 
parents: 
36419 
diff
changeset
 | 
88  | 
(axclasses, params, proven_classrels, proven_arities, inst_params, diff_classrels) =  | 
| 
36329
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
89  | 
  Data {axclasses = axclasses, params = params, proven_classrels = proven_classrels,
 | 
| 
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
90  | 
proven_arities = proven_arities, inst_params = inst_params,  | 
| 
36420
 
66fd989c8e15
tuned diff_classrels -- avoid slightly inefficient Symreltab.keys;
 
wenzelm 
parents: 
36419 
diff
changeset
 | 
91  | 
diff_classrels = diff_classrels};  | 
| 
 
66fd989c8e15
tuned diff_classrels -- avoid slightly inefficient Symreltab.keys;
 
wenzelm 
parents: 
36419 
diff
changeset
 | 
92  | 
|
| 
 
66fd989c8e15
tuned diff_classrels -- avoid slightly inefficient Symreltab.keys;
 
wenzelm 
parents: 
36419 
diff
changeset
 | 
93  | 
fun diff_table tab1 tab2 =  | 
| 
 
66fd989c8e15
tuned diff_classrels -- avoid slightly inefficient Symreltab.keys;
 
wenzelm 
parents: 
36419 
diff
changeset
 | 
94  | 
Symreltab.fold (fn (x, _) => if Symreltab.defined tab2 x then I else cons x) tab1 [];  | 
| 
36329
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
95  | 
|
| 36327 | 96  | 
structure Data = Theory_Data_PP  | 
| 22846 | 97  | 
(  | 
| 
36329
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
98  | 
type T = data;  | 
| 
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
99  | 
val empty =  | 
| 
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
100  | 
make_data (Symtab.empty, [], Symreltab.empty, Symtab.empty, (Symtab.empty, Symtab.empty), []);  | 
| 19574 | 101  | 
val extend = I;  | 
| 
36329
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
102  | 
fun merge pp  | 
| 
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
103  | 
      (Data {axclasses = axclasses1, params = params1, proven_classrels = proven_classrels1,
 | 
| 
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
104  | 
proven_arities = proven_arities1, inst_params = inst_params1,  | 
| 
36420
 
66fd989c8e15
tuned diff_classrels -- avoid slightly inefficient Symreltab.keys;
 
wenzelm 
parents: 
36419 
diff
changeset
 | 
105  | 
diff_classrels = diff_classrels1},  | 
| 
36329
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
106  | 
       Data {axclasses = axclasses2, params = params2, proven_classrels = proven_classrels2,
 | 
| 
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
107  | 
proven_arities = proven_arities2, inst_params = inst_params2,  | 
| 
36420
 
66fd989c8e15
tuned diff_classrels -- avoid slightly inefficient Symreltab.keys;
 
wenzelm 
parents: 
36419 
diff
changeset
 | 
108  | 
diff_classrels = diff_classrels2}) =  | 
| 36325 | 109  | 
let  | 
| 
36329
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
110  | 
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
 | 
111  | 
val params' =  | 
| 
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
112  | 
if null params1 then params2  | 
| 36417 | 113  | 
else fold_rev (fn p => if member (op =) params1 p then I else add_param pp p) params2 params1;  | 
| 
36329
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
114  | 
|
| 
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
115  | 
(*transitive closure of classrels and arity completion is done in Theory.at_begin hook*)  | 
| 
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
116  | 
val proven_classrels' = Symreltab.join (K #1) (proven_classrels1, proven_classrels2);  | 
| 
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
117  | 
val proven_arities' =  | 
| 
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
118  | 
Symtab.join (K (Library.merge (eq_fst op =))) (proven_arities1, proven_arities2);  | 
| 
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
119  | 
|
| 
36420
 
66fd989c8e15
tuned diff_classrels -- avoid slightly inefficient Symreltab.keys;
 
wenzelm 
parents: 
36419 
diff
changeset
 | 
120  | 
val diff_classrels' =  | 
| 
 
66fd989c8e15
tuned diff_classrels -- avoid slightly inefficient Symreltab.keys;
 
wenzelm 
parents: 
36419 
diff
changeset
 | 
121  | 
diff_table proven_classrels1 proven_classrels2 @  | 
| 
 
66fd989c8e15
tuned diff_classrels -- avoid slightly inefficient Symreltab.keys;
 
wenzelm 
parents: 
36419 
diff
changeset
 | 
122  | 
diff_table proven_classrels2 proven_classrels1 @  | 
| 
 
66fd989c8e15
tuned diff_classrels -- avoid slightly inefficient Symreltab.keys;
 
wenzelm 
parents: 
36419 
diff
changeset
 | 
123  | 
diff_classrels1 @ diff_classrels2;  | 
| 
36329
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
124  | 
|
| 
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
125  | 
val inst_params' =  | 
| 
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
126  | 
(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
 | 
127  | 
Symtab.merge (K true) (#2 inst_params1, #2 inst_params2));  | 
| 36325 | 128  | 
in  | 
| 
36420
 
66fd989c8e15
tuned diff_classrels -- avoid slightly inefficient Symreltab.keys;
 
wenzelm 
parents: 
36419 
diff
changeset
 | 
129  | 
make_data  | 
| 
 
66fd989c8e15
tuned diff_classrels -- avoid slightly inefficient Symreltab.keys;
 
wenzelm 
parents: 
36419 
diff
changeset
 | 
130  | 
(axclasses', params', proven_classrels', proven_arities', inst_params', diff_classrels')  | 
| 36325 | 131  | 
end;  | 
| 22846 | 132  | 
);  | 
| 6379 | 133  | 
|
| 
36329
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
134  | 
fun map_data f =  | 
| 
36420
 
66fd989c8e15
tuned diff_classrels -- avoid slightly inefficient Symreltab.keys;
 
wenzelm 
parents: 
36419 
diff
changeset
 | 
135  | 
  Data.map (fn Data {axclasses, params, proven_classrels, proven_arities, inst_params, diff_classrels} =>
 | 
| 
 
66fd989c8e15
tuned diff_classrels -- avoid slightly inefficient Symreltab.keys;
 
wenzelm 
parents: 
36419 
diff
changeset
 | 
136  | 
make_data (f (axclasses, params, proven_classrels, proven_arities, inst_params, diff_classrels)));  | 
| 
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_axclasses f =  | 
| 
36420
 
66fd989c8e15
tuned diff_classrels -- avoid slightly inefficient Symreltab.keys;
 
wenzelm 
parents: 
36419 
diff
changeset
 | 
139  | 
map_data (fn (axclasses, params, proven_classrels, proven_arities, inst_params, diff_classrels) =>  | 
| 
 
66fd989c8e15
tuned diff_classrels -- avoid slightly inefficient Symreltab.keys;
 
wenzelm 
parents: 
36419 
diff
changeset
 | 
140  | 
(f axclasses, params, proven_classrels, proven_arities, inst_params, diff_classrels));  | 
| 
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  | 
fun map_params f =  | 
| 
36420
 
66fd989c8e15
tuned diff_classrels -- avoid slightly inefficient Symreltab.keys;
 
wenzelm 
parents: 
36419 
diff
changeset
 | 
143  | 
map_data (fn (axclasses, params, proven_classrels, proven_arities, inst_params, diff_classrels) =>  | 
| 
 
66fd989c8e15
tuned diff_classrels -- avoid slightly inefficient Symreltab.keys;
 
wenzelm 
parents: 
36419 
diff
changeset
 | 
144  | 
(axclasses, f params, proven_classrels, proven_arities, inst_params, diff_classrels));  | 
| 
36329
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
145  | 
|
| 
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
146  | 
fun map_proven_classrels f =  | 
| 
36420
 
66fd989c8e15
tuned diff_classrels -- avoid slightly inefficient Symreltab.keys;
 
wenzelm 
parents: 
36419 
diff
changeset
 | 
147  | 
map_data (fn (axclasses, params, proven_classrels, proven_arities, inst_params, diff_classrels) =>  | 
| 
 
66fd989c8e15
tuned diff_classrels -- avoid slightly inefficient Symreltab.keys;
 
wenzelm 
parents: 
36419 
diff
changeset
 | 
148  | 
(axclasses, params, f proven_classrels, proven_arities, inst_params, diff_classrels));  | 
| 
36329
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
149  | 
|
| 
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
150  | 
fun map_proven_arities f =  | 
| 
36420
 
66fd989c8e15
tuned diff_classrels -- avoid slightly inefficient Symreltab.keys;
 
wenzelm 
parents: 
36419 
diff
changeset
 | 
151  | 
map_data (fn (axclasses, params, proven_classrels, proven_arities, inst_params, diff_classrels) =>  | 
| 
 
66fd989c8e15
tuned diff_classrels -- avoid slightly inefficient Symreltab.keys;
 
wenzelm 
parents: 
36419 
diff
changeset
 | 
152  | 
(axclasses, params, proven_classrels, f proven_arities, inst_params, diff_classrels));  | 
| 
36329
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
153  | 
|
| 
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
154  | 
fun map_inst_params f =  | 
| 
36420
 
66fd989c8e15
tuned diff_classrels -- avoid slightly inefficient Symreltab.keys;
 
wenzelm 
parents: 
36419 
diff
changeset
 | 
155  | 
map_data (fn (axclasses, params, proven_classrels, proven_arities, inst_params, diff_classrels) =>  | 
| 
 
66fd989c8e15
tuned diff_classrels -- avoid slightly inefficient Symreltab.keys;
 
wenzelm 
parents: 
36419 
diff
changeset
 | 
156  | 
(axclasses, params, proven_classrels, proven_arities, f inst_params, diff_classrels));  | 
| 
36329
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
157  | 
|
| 
36420
 
66fd989c8e15
tuned diff_classrels -- avoid slightly inefficient Symreltab.keys;
 
wenzelm 
parents: 
36419 
diff
changeset
 | 
158  | 
val clear_diff_classrels =  | 
| 
36329
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
159  | 
map_data (fn (axclasses, params, proven_classrels, proven_arities, inst_params, _) =>  | 
| 
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
160  | 
(axclasses, params, proven_classrels, proven_arities, inst_params, []));  | 
| 
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
161  | 
|
| 
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
162  | 
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
 | 
163  | 
|
| 
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
164  | 
val axclasses_of = #axclasses o rep_data;  | 
| 
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
165  | 
val params_of = #params o rep_data;  | 
| 
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
166  | 
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
 | 
167  | 
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
 | 
168  | 
val inst_params_of = #inst_params o rep_data;  | 
| 
36420
 
66fd989c8e15
tuned diff_classrels -- avoid slightly inefficient Symreltab.keys;
 
wenzelm 
parents: 
36419 
diff
changeset
 | 
169  | 
val diff_classrels_of = #diff_classrels o rep_data;  | 
| 
36329
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
170  | 
|
| 6379 | 171  | 
|
| 36346 | 172  | 
(* axclasses with parameters *)  | 
| 19392 | 173  | 
|
| 
24929
 
408becab067e
renamed AxClass.get_definition to AxClass.get_info (again);
 
wenzelm 
parents: 
24920 
diff
changeset
 | 
174  | 
fun get_info thy c =  | 
| 
36329
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
175  | 
(case Symtab.lookup (axclasses_of thy) c of  | 
| 36327 | 176  | 
SOME info => info  | 
| 21919 | 177  | 
  | NONE => error ("No such axclass: " ^ quote c));
 | 
| 6379 | 178  | 
|
| 36327 | 179  | 
fun all_params_of thy S =  | 
| 
36329
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
180  | 
let val params = params_of thy;  | 
| 36327 | 181  | 
in fold (fn (x, c) => if Sign.subsort thy (S, [c]) then cons x else I) params [] end;  | 
| 19460 | 182  | 
|
| 
36329
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
183  | 
fun class_of_param thy = AList.lookup (op =) (params_of thy);  | 
| 19460 | 184  | 
|
| 21919 | 185  | 
|
| 19511 | 186  | 
(* maintain instances *)  | 
| 19503 | 187  | 
|
| 
36329
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
188  | 
val classrel_prefix = "classrel_";  | 
| 
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
189  | 
val arity_prefix = "arity_";  | 
| 25486 | 190  | 
|
| 
36329
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
191  | 
fun instance_name (a, c) = Long_Name.base_name c ^ "_" ^ Long_Name.base_name a;  | 
| 19528 | 192  | 
|
193  | 
||
| 
36421
 
066e35d1c0d7
tuned classrel completion -- bypass composition with reflexive edges;
 
wenzelm 
parents: 
36420 
diff
changeset
 | 
194  | 
infix 0 RSO;  | 
| 
 
066e35d1c0d7
tuned classrel completion -- bypass composition with reflexive edges;
 
wenzelm 
parents: 
36420 
diff
changeset
 | 
195  | 
|
| 36456 | 196  | 
fun (SOME a) RSO (SOME b) = SOME (a RS b)  | 
197  | 
| x RSO NONE = x  | 
|
198  | 
| NONE RSO y = y;  | 
|
| 
36421
 
066e35d1c0d7
tuned classrel completion -- bypass composition with reflexive edges;
 
wenzelm 
parents: 
36420 
diff
changeset
 | 
199  | 
|
| 19574 | 200  | 
fun the_classrel thy (c1, c2) =  | 
| 
36329
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
201  | 
(case Symreltab.lookup (proven_classrels_of thy) (c1, c2) of  | 
| 
36881
 
4de023c28a84
the_classrel/the_arity: avoid Thm.transfer for proofterm version -- theory might already have become stale within the proof_body future;
 
wenzelm 
parents: 
36768 
diff
changeset
 | 
202  | 
SOME thm => thm  | 
| 24920 | 203  | 
  | NONE => error ("Unproven class relation " ^
 | 
| 42360 | 204  | 
Syntax.string_of_classrel (Proof_Context.init_global thy) [c1, c2])); (* FIXME stale thy (!?) *)  | 
| 19574 | 205  | 
|
| 36325 | 206  | 
fun put_trancl_classrel ((c1, c2), th) thy =  | 
207  | 
let  | 
|
| 
36328
 
4d9deabf6474
replaced Sorts.rep_algebra by slightly more abstract selectors classes_of/arities_of;
 
wenzelm 
parents: 
36327 
diff
changeset
 | 
208  | 
val classes = Sorts.classes_of (Sign.classes_of thy);  | 
| 
36329
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
209  | 
val classrels = proven_classrels_of thy;  | 
| 36325 | 210  | 
|
211  | 
fun reflcl_classrel (c1', c2') =  | 
|
| 
36881
 
4de023c28a84
the_classrel/the_arity: avoid Thm.transfer for proofterm version -- theory might already have become stale within the proof_body future;
 
wenzelm 
parents: 
36768 
diff
changeset
 | 
212  | 
if c1' = c2' then NONE else SOME (Thm.transfer thy (the_classrel thy (c1', c2')));  | 
| 36325 | 213  | 
fun gen_classrel (c1_pred, c2_succ) =  | 
214  | 
let  | 
|
| 
36421
 
066e35d1c0d7
tuned classrel completion -- bypass composition with reflexive edges;
 
wenzelm 
parents: 
36420 
diff
changeset
 | 
215  | 
val th' =  | 
| 
 
066e35d1c0d7
tuned classrel completion -- bypass composition with reflexive edges;
 
wenzelm 
parents: 
36420 
diff
changeset
 | 
216  | 
the ((reflcl_classrel (c1_pred, c1) RSO SOME th) RSO reflcl_classrel (c2, c2_succ))  | 
| 
 
066e35d1c0d7
tuned classrel completion -- bypass composition with reflexive edges;
 
wenzelm 
parents: 
36420 
diff
changeset
 | 
217  | 
|> Drule.instantiate' [SOME (ctyp_of thy (TVar ((Name.aT, 0), [])))] []  | 
| 36327 | 218  | 
|> 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
 | 
219  | 
in ((c1_pred, c2_succ), th') end;  | 
| 36325 | 220  | 
|
| 36327 | 221  | 
val new_classrels =  | 
222  | 
Library.map_product pair (c1 :: Graph.imm_preds classes c1) (c2 :: Graph.imm_succs classes c2)  | 
|
| 
36421
 
066e35d1c0d7
tuned classrel completion -- bypass composition with reflexive edges;
 
wenzelm 
parents: 
36420 
diff
changeset
 | 
223  | 
|> filter_out ((op =) orf Symreltab.defined classrels)  | 
| 36327 | 224  | 
|> map gen_classrel;  | 
225  | 
val needed = not (null new_classrels);  | 
|
| 36325 | 226  | 
in  | 
227  | 
(needed,  | 
|
| 
36329
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
228  | 
if needed then map_proven_classrels (fold Symreltab.update new_classrels) thy  | 
| 
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
229  | 
else thy)  | 
| 36325 | 230  | 
end;  | 
231  | 
||
232  | 
fun complete_classrels thy =  | 
|
233  | 
let  | 
|
| 
36329
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
234  | 
val classrels = proven_classrels_of thy;  | 
| 
36420
 
66fd989c8e15
tuned diff_classrels -- avoid slightly inefficient Symreltab.keys;
 
wenzelm 
parents: 
36419 
diff
changeset
 | 
235  | 
val diff_classrels = diff_classrels_of thy;  | 
| 36325 | 236  | 
val (needed, thy') = (false, thy) |>  | 
| 36417 | 237  | 
fold (fn rel => fn (needed, thy) =>  | 
| 
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
 | 
238  | 
put_trancl_classrel (rel, Symreltab.lookup classrels rel |> the) thy  | 
| 36325 | 239  | 
|>> (fn b => needed orelse b))  | 
| 
36420
 
66fd989c8e15
tuned diff_classrels -- avoid slightly inefficient Symreltab.keys;
 
wenzelm 
parents: 
36419 
diff
changeset
 | 
240  | 
diff_classrels;  | 
| 36325 | 241  | 
in  | 
| 
36420
 
66fd989c8e15
tuned diff_classrels -- avoid slightly inefficient Symreltab.keys;
 
wenzelm 
parents: 
36419 
diff
changeset
 | 
242  | 
if null diff_classrels then NONE  | 
| 
 
66fd989c8e15
tuned diff_classrels -- avoid slightly inefficient Symreltab.keys;
 
wenzelm 
parents: 
36419 
diff
changeset
 | 
243  | 
else SOME (clear_diff_classrels thy')  | 
| 36325 | 244  | 
end;  | 
| 19503 | 245  | 
|
246  | 
||
| 36740 | 247  | 
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
 | 
248  | 
(case AList.lookup (op =) (Symtab.lookup_list (proven_arities_of thy) a) (c, Ss) of  | 
| 
36881
 
4de023c28a84
the_classrel/the_arity: avoid Thm.transfer for proofterm version -- theory might already have become stale within the proof_body future;
 
wenzelm 
parents: 
36768 
diff
changeset
 | 
249  | 
SOME (thm, _) => thm  | 
| 24920 | 250  | 
  | NONE => error ("Unproven type arity " ^
 | 
| 42360 | 251  | 
Syntax.string_of_arity (Proof_Context.init_global thy) (a, Ss, [c]))); (* FIXME stale thy (!?) *)  | 
| 19503 | 252  | 
|
| 33172 | 253  | 
fun thynames_of_arity thy (c, a) =  | 
| 
36329
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
254  | 
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
 | 
255  | 
|> map_filter (fn ((c', _), (_, name)) => if c = c' then SOME name else NONE)  | 
| 27497 | 256  | 
|> rev;  | 
| 27397 | 257  | 
|
| 
36419
 
7aea10d10113
tuned aritiy completion -- slightly less intermediate data structures;
 
wenzelm 
parents: 
36418 
diff
changeset
 | 
258  | 
fun insert_arity_completions thy t ((c, Ss), ((th, thy_name))) (finished, arities) =  | 
| 27547 | 259  | 
let  | 
260  | 
val algebra = Sign.classes_of thy;  | 
|
| 36417 | 261  | 
val ars = Symtab.lookup_list arities t;  | 
| 33172 | 262  | 
val super_class_completions =  | 
263  | 
Sign.super_classes thy c  | 
|
| 36417 | 264  | 
|> filter_out (fn c1 => exists (fn ((c2, Ss2), _) =>  | 
265  | 
c1 = c2 andalso Sorts.sorts_le algebra (Ss2, Ss)) ars);  | 
|
266  | 
||
267  | 
val names = Name.invents Name.context Name.aT (length Ss);  | 
|
268  | 
val std_vars = map (fn a => SOME (ctyp_of thy (TVar ((a, 0), [])))) names;  | 
|
269  | 
||
| 36325 | 270  | 
val completions = super_class_completions |> map (fn c1 =>  | 
271  | 
let  | 
|
| 
36419
 
7aea10d10113
tuned aritiy completion -- slightly less intermediate data structures;
 
wenzelm 
parents: 
36418 
diff
changeset
 | 
272  | 
val th1 =  | 
| 
36881
 
4de023c28a84
the_classrel/the_arity: avoid Thm.transfer for proofterm version -- theory might already have become stale within the proof_body future;
 
wenzelm 
parents: 
36768 
diff
changeset
 | 
273  | 
(th RS Thm.transfer thy (the_classrel thy (c, c1)))  | 
| 36417 | 274  | 
|> Drule.instantiate' std_vars []  | 
| 36327 | 275  | 
|> 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
 | 
276  | 
in ((th1, thy_name), c1) end);  | 
| 
36419
 
7aea10d10113
tuned aritiy completion -- slightly less intermediate data structures;
 
wenzelm 
parents: 
36418 
diff
changeset
 | 
277  | 
|
| 
 
7aea10d10113
tuned aritiy completion -- slightly less intermediate data structures;
 
wenzelm 
parents: 
36418 
diff
changeset
 | 
278  | 
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
 | 
279  | 
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
 | 
280  | 
in (finished', arities') end;  | 
| 19503 | 281  | 
|
| 27547 | 282  | 
fun put_arity ((t, Ss, c), th) thy =  | 
| 
36419
 
7aea10d10113
tuned aritiy completion -- slightly less intermediate data structures;
 
wenzelm 
parents: 
36418 
diff
changeset
 | 
283  | 
let val ar = ((c, Ss), (th, Context.theory_name thy)) in  | 
| 27547 | 284  | 
thy  | 
| 
36329
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
285  | 
|> map_proven_arities  | 
| 
36419
 
7aea10d10113
tuned aritiy completion -- slightly less intermediate data structures;
 
wenzelm 
parents: 
36418 
diff
changeset
 | 
286  | 
(Symtab.insert_list (eq_fst op =) (t, ar) #>  | 
| 
 
7aea10d10113
tuned aritiy completion -- slightly less intermediate data structures;
 
wenzelm 
parents: 
36418 
diff
changeset
 | 
287  | 
curry (insert_arity_completions thy t ar) true #> #2)  | 
| 27547 | 288  | 
end;  | 
| 19503 | 289  | 
|
| 31944 | 290  | 
fun complete_arities thy =  | 
| 27547 | 291  | 
let  | 
| 
36329
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
292  | 
val arities = proven_arities_of thy;  | 
| 
36419
 
7aea10d10113
tuned aritiy completion -- slightly less intermediate data structures;
 
wenzelm 
parents: 
36418 
diff
changeset
 | 
293  | 
val (finished, arities') =  | 
| 
 
7aea10d10113
tuned aritiy completion -- slightly less intermediate data structures;
 
wenzelm 
parents: 
36418 
diff
changeset
 | 
294  | 
Symtab.fold (fn (t, ars) => fold (insert_arity_completions thy t) ars) arities (true, arities);  | 
| 33172 | 295  | 
in  | 
| 
36419
 
7aea10d10113
tuned aritiy completion -- slightly less intermediate data structures;
 
wenzelm 
parents: 
36418 
diff
changeset
 | 
296  | 
if finished then NONE  | 
| 
36329
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
297  | 
else SOME (map_proven_arities (K arities') thy)  | 
| 27547 | 298  | 
end;  | 
299  | 
||
| 36325 | 300  | 
val _ = Context.>> (Context.map_theory  | 
| 36740 | 301  | 
(Theory.at_begin complete_classrels #>  | 
302  | 
Theory.at_begin complete_arities));  | 
|
| 27397 | 303  | 
|
| 36740 | 304  | 
val _ = Proofterm.install_axclass_proofs  | 
305  | 
  {classrel_proof = Thm.proof_of oo the_classrel,
 | 
|
306  | 
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
 | 
307  | 
|
| 27397 | 308  | 
|
| 
25597
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
309  | 
(* maintain instance parameters *)  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
310  | 
|
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
311  | 
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
 | 
312  | 
(case Symtab.lookup (the_default Symtab.empty (Symtab.lookup (#1 (inst_params_of thy)) c)) tyco of  | 
| 36327 | 313  | 
SOME c' => c'  | 
314  | 
  | 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
 | 
315  | 
|
| 36327 | 316  | 
fun add_inst_param (c, tyco) inst =  | 
317  | 
(map_inst_params o apfst o Symtab.map_default (c, Symtab.empty)) (Symtab.update_new (tyco, inst))  | 
|
| 36417 | 318  | 
#> (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
 | 
319  | 
|
| 
36329
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
320  | 
val inst_of_param = Symtab.lookup o #2 o inst_params_of;  | 
| 36417 | 321  | 
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
 | 
322  | 
|
| 36327 | 323  | 
fun inst_thms thy =  | 
| 
36329
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
324  | 
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
 | 
325  | 
|
| 36417 | 326  | 
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
 | 
327  | 
|
| 
41228
 
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
 
wenzelm 
parents: 
39557 
diff
changeset
 | 
328  | 
fun unoverload thy = Raw_Simplifier.simplify true (inst_thms thy);  | 
| 
 
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
 
wenzelm 
parents: 
39557 
diff
changeset
 | 
329  | 
fun overload thy = Raw_Simplifier.simplify true (map Thm.symmetric (inst_thms thy));  | 
| 
25597
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
330  | 
|
| 
41228
 
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
 
wenzelm 
parents: 
39557 
diff
changeset
 | 
331  | 
fun unoverload_conv thy = Raw_Simplifier.rewrite true (inst_thms thy);  | 
| 
 
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
 
wenzelm 
parents: 
39557 
diff
changeset
 | 
332  | 
fun overload_conv thy = Raw_Simplifier.rewrite true (map Thm.symmetric (inst_thms thy));  | 
| 
25597
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
333  | 
|
| 36327 | 334  | 
fun lookup_inst_param consts params (c, T) =  | 
335  | 
(case get_inst_tyco consts (c, T) of  | 
|
336  | 
SOME tyco => AList.lookup (op =) params (c, tyco)  | 
|
337  | 
| NONE => NONE);  | 
|
| 31249 | 338  | 
|
| 
25597
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
339  | 
fun unoverload_const thy (c_ty as (c, _)) =  | 
| 36327 | 340  | 
if is_some (class_of_param thy c) then  | 
341  | 
(case get_inst_tyco (Sign.consts_of thy) c_ty of  | 
|
342  | 
SOME tyco => try (param_of_inst thy) (c, tyco) |> the_default c  | 
|
343  | 
| NONE => c)  | 
|
| 33969 | 344  | 
else c;  | 
| 
25597
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
345  | 
|
| 
404
 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 
wenzelm 
parents:  
diff
changeset
 | 
346  | 
|
| 36327 | 347  | 
|
| 19511 | 348  | 
(** instances **)  | 
| 
19418
 
03b01c9314fc
reworded add_axclass(_i): canonical specification format,
 
wenzelm 
parents: 
19405 
diff
changeset
 | 
349  | 
|
| 
 
03b01c9314fc
reworded add_axclass(_i): canonical specification format,
 
wenzelm 
parents: 
19405 
diff
changeset
 | 
350  | 
(* class relations *)  | 
| 
 
03b01c9314fc
reworded add_axclass(_i): canonical specification format,
 
wenzelm 
parents: 
19405 
diff
changeset
 | 
351  | 
|
| 19405 | 352  | 
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
 | 
353  | 
let  | 
| 
26939
 
1035c89b4c02
moved global pretty/string_of functions from Sign to Syntax;
 
wenzelm 
parents: 
26516 
diff
changeset
 | 
354  | 
val string_of_sort = Syntax.string_of_sort_global thy;  | 
| 19405 | 355  | 
val (c1, c2) = pairself (Sign.certify_class thy) raw_rel;  | 
| 24271 | 356  | 
val _ = Sign.primitive_classrel (c1, c2) (Theory.copy thy);  | 
| 19405 | 357  | 
val _ =  | 
| 19460 | 358  | 
(case subtract (op =) (all_params_of thy [c1]) (all_params_of thy [c2]) of  | 
| 19405 | 359  | 
[] => ()  | 
| 
26939
 
1035c89b4c02
moved global pretty/string_of functions from Sign to Syntax;
 
wenzelm 
parents: 
26516 
diff
changeset
 | 
360  | 
      | 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
 | 
361  | 
commas_quote xs ^ " of " ^ string_of_sort [c2], [], []));  | 
| 19405 | 362  | 
in (c1, c2) end;  | 
363  | 
||
364  | 
fun read_classrel thy raw_rel =  | 
|
| 42360 | 365  | 
cert_classrel thy (pairself (Proof_Context.read_class (Proof_Context.init_global thy)) raw_rel)  | 
| 19405 | 366  | 
handle TYPE (msg, _, _) => error msg;  | 
367  | 
||
368  | 
||
| 
25597
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
369  | 
(* declaration and definition of instances of overloaded constants *)  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
370  | 
|
| 
35201
 
c2ddb9395436
axclass: more precise treatment of naming vs. binding;
 
wenzelm 
parents: 
35021 
diff
changeset
 | 
371  | 
fun inst_tyco_of thy (c, T) =  | 
| 
 
c2ddb9395436
axclass: more precise treatment of naming vs. binding;
 
wenzelm 
parents: 
35021 
diff
changeset
 | 
372  | 
(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
 | 
373  | 
SOME tyco => tyco  | 
| 
 
c2ddb9395436
axclass: more precise treatment of naming vs. binding;
 
wenzelm 
parents: 
35021 
diff
changeset
 | 
374  | 
  | NONE => error ("Illegal type for instantiation of class parameter: " ^
 | 
| 
 
c2ddb9395436
axclass: more precise treatment of naming vs. binding;
 
wenzelm 
parents: 
35021 
diff
changeset
 | 
375  | 
quote (c ^ " :: " ^ Syntax.string_of_typ_global thy T)));  | 
| 31249 | 376  | 
|
| 
25597
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
377  | 
fun declare_overloaded (c, T) thy =  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
378  | 
let  | 
| 
35201
 
c2ddb9395436
axclass: more precise treatment of naming vs. binding;
 
wenzelm 
parents: 
35021 
diff
changeset
 | 
379  | 
val class =  | 
| 
 
c2ddb9395436
axclass: more precise treatment of naming vs. binding;
 
wenzelm 
parents: 
35021 
diff
changeset
 | 
380  | 
(case class_of_param thy c of  | 
| 
 
c2ddb9395436
axclass: more precise treatment of naming vs. binding;
 
wenzelm 
parents: 
35021 
diff
changeset
 | 
381  | 
SOME class => class  | 
| 
 
c2ddb9395436
axclass: more precise treatment of naming vs. binding;
 
wenzelm 
parents: 
35021 
diff
changeset
 | 
382  | 
      | NONE => error ("Not a class parameter: " ^ quote c));
 | 
| 31249 | 383  | 
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
 | 
384  | 
val name_inst = instance_name (tyco, class) ^ "_inst";  | 
| 36417 | 385  | 
val c' = instance_name (tyco, c);  | 
| 
25597
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
386  | 
val T' = Type.strip_sorts T;  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
387  | 
in  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
388  | 
thy  | 
| 
35201
 
c2ddb9395436
axclass: more precise treatment of naming vs. binding;
 
wenzelm 
parents: 
35021 
diff
changeset
 | 
389  | 
|> Sign.qualified_path true (Binding.name name_inst)  | 
| 
42375
 
774df7c59508
report Name_Space.declare/define, relatively to context;
 
wenzelm 
parents: 
42360 
diff
changeset
 | 
390  | 
|> 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
 | 
391  | 
|-> (fn const' as Const (c'', _) =>  | 
| 
42375
 
774df7c59508
report Name_Space.declare/define, relatively to context;
 
wenzelm 
parents: 
42360 
diff
changeset
 | 
392  | 
Thm.add_def_global false true  | 
| 
30344
 
10a67c5ddddb
more uniform handling of binding in targets and derived elements;
 
wenzelm 
parents: 
30280 
diff
changeset
 | 
393  | 
(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
 | 
394  | 
#>> apsnd Thm.varifyT_global  | 
| 
 
19deea200358
Thm.add_axiom/add_def: return internal name of foundational axiom;
 
wenzelm 
parents: 
35961 
diff
changeset
 | 
395  | 
#-> (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
 | 
396  | 
#> Global_Theory.add_thm ((Binding.conceal (Binding.name c'), thm), [])  | 
| 36417 | 397  | 
#> #2  | 
| 
36106
 
19deea200358
Thm.add_axiom/add_def: return internal name of foundational axiom;
 
wenzelm 
parents: 
35961 
diff
changeset
 | 
398  | 
#> pair (Const (c, T))))  | 
| 
35201
 
c2ddb9395436
axclass: more precise treatment of naming vs. binding;
 
wenzelm 
parents: 
35021 
diff
changeset
 | 
399  | 
||> Sign.restore_naming thy  | 
| 
25597
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
400  | 
end;  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
401  | 
|
| 
30519
 
c05c0199826f
coherent binding policy with primitive target operations
 
haftmann 
parents: 
30469 
diff
changeset
 | 
402  | 
fun define_overloaded b (c, t) thy =  | 
| 
25597
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
403  | 
let  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
404  | 
val T = Term.fastype_of t;  | 
| 31249 | 405  | 
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
 | 
406  | 
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
 | 
407  | 
val prop = Logic.mk_equals (Const (c', T), t);  | 
| 36417 | 408  | 
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
 | 
409  | 
in  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
410  | 
thy  | 
| 
42375
 
774df7c59508
report Name_Space.declare/define, relatively to context;
 
wenzelm 
parents: 
42360 
diff
changeset
 | 
411  | 
|> Thm.add_def_global false false (b', prop)  | 
| 38383 | 412  | 
|>> (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
 | 
413  | 
end;  | 
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
414  | 
|
| 
 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 
haftmann 
parents: 
25486 
diff
changeset
 | 
415  | 
|
| 
30951
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
416  | 
(* primitive rules *)  | 
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
417  | 
|
| 37249 | 418  | 
fun add_classrel raw_th thy =  | 
| 
30951
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
419  | 
let  | 
| 
31948
 
ea8c8bf47ce3
add_classrel/arity: strip_shyps of stored result;
 
wenzelm 
parents: 
31944 
diff
changeset
 | 
420  | 
val th = Thm.strip_shyps (Thm.transfer thy raw_th);  | 
| 
 
ea8c8bf47ce3
add_classrel/arity: strip_shyps of stored result;
 
wenzelm 
parents: 
31944 
diff
changeset
 | 
421  | 
val prop = Thm.plain_prop_of th;  | 
| 
30951
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
422  | 
    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
 | 
423  | 
val rel = Logic.dest_classrel prop handle TERM _ => err ();  | 
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
424  | 
val (c1, c2) = cert_classrel thy rel handle TYPE _ => err ();  | 
| 
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
 | 
425  | 
val (th', thy') = Global_Theory.store_thm  | 
| 37249 | 426  | 
(Binding.name (prefix classrel_prefix (Logic.name_classrel (c1, c2))), th) thy;  | 
| 
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  | 
| 
37232
 
c10fb22a5e0c
classrel and arity theorems are now stored under proper name in theory. add_arity and
 
berghofe 
parents: 
36935 
diff
changeset
 | 
429  | 
|> 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
 | 
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'  | 
| 
30951
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
432  | 
|> Sign.primitive_classrel (c1, c2)  | 
| 
37232
 
c10fb22a5e0c
classrel and arity theorems are now stored under proper name in theory. add_arity and
 
berghofe 
parents: 
36935 
diff
changeset
 | 
433  | 
|> (#2 oo put_trancl_classrel) ((c1, c2), th'')  | 
| 
30951
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
434  | 
|> perhaps complete_arities  | 
| 
 
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  | 
|
| 37249 | 437  | 
fun add_arity raw_th thy =  | 
| 
30951
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
438  | 
let  | 
| 
31948
 
ea8c8bf47ce3
add_classrel/arity: strip_shyps of stored result;
 
wenzelm 
parents: 
31944 
diff
changeset
 | 
439  | 
val th = Thm.strip_shyps (Thm.transfer thy raw_th);  | 
| 
 
ea8c8bf47ce3
add_classrel/arity: strip_shyps of stored result;
 
wenzelm 
parents: 
31944 
diff
changeset
 | 
440  | 
val prop = Thm.plain_prop_of th;  | 
| 
30951
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
441  | 
    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
 | 
442  | 
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
 | 
443  | 
|
| 
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
 | 
444  | 
val (th', thy') = Global_Theory.store_thm  | 
| 37249 | 445  | 
(Binding.name (prefix arity_prefix (Logic.name_arity arity)), th) thy;  | 
| 36417 | 446  | 
|
447  | 
val args = Name.names Name.context Name.aT Ss;  | 
|
448  | 
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
 | 
449  | 
val std_vars = map (fn (a, S) => SOME (ctyp_of thy' (TVar ((a, 0), [])))) args;  | 
| 36417 | 450  | 
|
| 
37232
 
c10fb22a5e0c
classrel and arity theorems are now stored under proper name in theory. add_arity and
 
berghofe 
parents: 
36935 
diff
changeset
 | 
451  | 
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
 | 
452  | 
|> 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
 | 
453  | 
|> 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
 | 
454  | 
|> (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
 | 
455  | 
val th'' = th'  | 
| 
36934
 
ae0809cff6f0
normalize atyp names after unconstrainT, which may rename atyps arbitrarily;
 
krauss 
parents: 
36881 
diff
changeset
 | 
456  | 
|> Thm.unconstrainT  | 
| 
 
ae0809cff6f0
normalize atyp names after unconstrainT, which may rename atyps arbitrarily;
 
krauss 
parents: 
36881 
diff
changeset
 | 
457  | 
|> Drule.instantiate' std_vars [];  | 
| 
30951
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
458  | 
in  | 
| 
37232
 
c10fb22a5e0c
classrel and arity theorems are now stored under proper name in theory. add_arity and
 
berghofe 
parents: 
36935 
diff
changeset
 | 
459  | 
thy'  | 
| 36417 | 460  | 
|> fold (#2 oo declare_overloaded) missing_params  | 
| 
30951
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
461  | 
|> Sign.primitive_arity (t, Ss, [c])  | 
| 
37232
 
c10fb22a5e0c
classrel and arity theorems are now stored under proper name in theory. add_arity and
 
berghofe 
parents: 
36935 
diff
changeset
 | 
462  | 
|> put_arity ((t, Ss, c), th'')  | 
| 
30951
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
463  | 
end;  | 
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
464  | 
|
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
465  | 
|
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
466  | 
(* tactical proofs *)  | 
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
467  | 
|
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
468  | 
fun prove_classrel raw_rel tac thy =  | 
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
469  | 
let  | 
| 42360 | 470  | 
val ctxt = Proof_Context.init_global thy;  | 
| 
30951
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
471  | 
val (c1, c2) = cert_classrel thy raw_rel;  | 
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
472  | 
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
 | 
473  | 
      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
 | 
474  | 
quote (Syntax.string_of_classrel ctxt [c1, c2]));  | 
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
475  | 
in  | 
| 37249 | 476  | 
thy |> add_classrel th  | 
| 
30951
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
477  | 
end;  | 
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
478  | 
|
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
479  | 
fun prove_arity raw_arity tac thy =  | 
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
480  | 
let  | 
| 42360 | 481  | 
val ctxt = Proof_Context.init_global thy;  | 
482  | 
val arity = Proof_Context.cert_arity ctxt raw_arity;  | 
|
| 
30951
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
483  | 
val names = map (prefix arity_prefix) (Logic.name_arities arity);  | 
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
484  | 
val props = Logic.mk_arities arity;  | 
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
485  | 
val ths = Goal.prove_multi ctxt [] [] props  | 
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
486  | 
(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
 | 
487  | 
        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
 | 
488  | 
quote (Syntax.string_of_arity ctxt arity));  | 
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
489  | 
in  | 
| 37249 | 490  | 
thy |> fold add_arity ths  | 
| 
30951
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
491  | 
end;  | 
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
492  | 
|
| 
 
a6e26a248f03
formal declaration of undefined parameters after class instantiation
 
haftmann 
parents: 
30519 
diff
changeset
 | 
493  | 
|
| 19511 | 494  | 
|
495  | 
(** class definitions **)  | 
|
496  | 
||
| 23421 | 497  | 
fun split_defined n eq =  | 
498  | 
let  | 
|
499  | 
val intro =  | 
|
500  | 
(eq RS Drule.equal_elim_rule2)  | 
|
501  | 
|> Conjunction.curry_balanced n  | 
|
502  | 
|> n = 0 ? Thm.eq_assumption 1;  | 
|
503  | 
val dests =  | 
|
504  | 
if n = 0 then []  | 
|
505  | 
else  | 
|
506  | 
(eq RS Drule.equal_elim_rule1)  | 
|
| 32765 | 507  | 
|> Balanced_Tree.dest (fn th =>  | 
| 23421 | 508  | 
(th RS Conjunction.conjunctionD1, th RS Conjunction.conjunctionD2)) n;  | 
509  | 
in (intro, dests) end;  | 
|
510  | 
||
| 24964 | 511  | 
fun define_class (bclass, raw_super) raw_params raw_specs thy =  | 
| 19511 | 512  | 
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
 | 
513  | 
val ctxt = Syntax.init_pretty_global thy;  | 
| 19511 | 514  | 
|
515  | 
||
| 24964 | 516  | 
(* class *)  | 
| 19511 | 517  | 
|
| 
30344
 
10a67c5ddddb
more uniform handling of binding in targets and derived elements;
 
wenzelm 
parents: 
30280 
diff
changeset
 | 
518  | 
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
 | 
519  | 
val class = Sign.full_name thy bclass;  | 
| 24731 | 520  | 
val super = Sign.minimize_sort thy (Sign.certify_sort thy raw_super);  | 
| 19511 | 521  | 
|
| 24964 | 522  | 
fun check_constraint (a, S) =  | 
523  | 
if Sign.subsort thy (super, S) then ()  | 
|
524  | 
      else error ("Sort constraint of type variable " ^
 | 
|
| 
39134
 
917b4b6ba3d2
turned show_sorts/show_types into proper configuration options;
 
wenzelm 
parents: 
39133 
diff
changeset
 | 
525  | 
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
 | 
526  | 
" needs to be weaker than " ^ Syntax.string_of_sort ctxt super);  | 
| 24964 | 527  | 
|
528  | 
||
529  | 
(* params *)  | 
|
530  | 
||
531  | 
val params = raw_params |> map (fn p =>  | 
|
532  | 
let  | 
|
533  | 
val T = Sign.the_const_type thy p;  | 
|
534  | 
val _ =  | 
|
535  | 
(case Term.add_tvarsT T [] of  | 
|
536  | 
[((a, _), S)] => check_constraint (a, S)  | 
|
537  | 
          | _ => error ("Exactly one type variable expected in class parameter " ^ quote p));
 | 
|
538  | 
val T' = Term.map_type_tvar (fn _ => TFree (Name.aT, [class])) T;  | 
|
539  | 
in (p, T') end);  | 
|
540  | 
||
541  | 
||
542  | 
(* axioms *)  | 
|
543  | 
||
| 19511 | 544  | 
fun prep_axiom t =  | 
545  | 
(case Term.add_tfrees t [] of  | 
|
| 24964 | 546  | 
[(a, S)] => check_constraint (a, S)  | 
547  | 
| [] => ()  | 
|
| 
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
 | 
548  | 
      | _ => error ("Multiple type variables in class axiom:\n" ^ Syntax.string_of_term ctxt t);
 | 
| 24964 | 549  | 
t  | 
550  | 
|> Term.map_types (Term.map_atyps (fn TFree _ => Term.aT [] | U => U))  | 
|
551  | 
|> Logic.close_form);  | 
|
| 19511 | 552  | 
|
| 24681 | 553  | 
val axiomss = map (map (prep_axiom o Sign.cert_prop thy) o snd) raw_specs;  | 
| 22745 | 554  | 
val name_atts = map fst raw_specs;  | 
| 19511 | 555  | 
|
556  | 
||
557  | 
(* definition *)  | 
|
558  | 
||
| 35854 | 559  | 
val conjs = Logic.mk_of_sort (Term.aT [], super) @ flat axiomss;  | 
| 19511 | 560  | 
val class_eq =  | 
| 
31943
 
5e960a0780a2
renamed inclass/Inclass to of_class/OfClass, in accordance to of_sort;
 
wenzelm 
parents: 
31904 
diff
changeset
 | 
561  | 
Logic.mk_equals (Logic.mk_of_class (Term.aT [], class), Logic.mk_conjunction_balanced conjs);  | 
| 19511 | 562  | 
|
563  | 
val ([def], def_thy) =  | 
|
564  | 
thy  | 
|
565  | 
|> 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
 | 
566  | 
|> Global_Theory.add_defs false [((Thm.def_binding bconst, class_eq), [])];  | 
| 19511 | 567  | 
val (raw_intro, (raw_classrel, raw_axioms)) =  | 
| 23421 | 568  | 
split_defined (length conjs) def ||> chop (length super);  | 
| 19392 | 569  | 
|
| 
19418
 
03b01c9314fc
reworded add_axclass(_i): canonical specification format,
 
wenzelm 
parents: 
19405 
diff
changeset
 | 
570  | 
|
| 19511 | 571  | 
(* facts *)  | 
572  | 
||
573  | 
val class_triv = Thm.class_triv def_thy class;  | 
|
574  | 
val ([(_, [intro]), (_, classrel), (_, axioms)], facts_thy) =  | 
|
575  | 
def_thy  | 
|
| 
35201
 
c2ddb9395436
axclass: more precise treatment of naming vs. binding;
 
wenzelm 
parents: 
35021 
diff
changeset
 | 
576  | 
|> 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
 | 
577  | 
|> Global_Theory.note_thmss ""  | 
| 36326 | 578  | 
[((Binding.name "intro", []), [([Drule.export_without_context raw_intro], [])]),  | 
579  | 
((Binding.name "super", []), [(map Drule.export_without_context raw_classrel, [])]),  | 
|
580  | 
((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
 | 
581  | 
[(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
 | 
582  | 
||> Sign.restore_naming def_thy;  | 
| 19511 | 583  | 
|
| 24847 | 584  | 
|
| 19511 | 585  | 
(* result *)  | 
586  | 
||
| 
36329
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
587  | 
val axclass = make_axclass (def, intro, axioms, params);  | 
| 19511 | 588  | 
val result_thy =  | 
589  | 
facts_thy  | 
|
| 36417 | 590  | 
|> fold (#2 oo put_trancl_classrel) (map (pair class) super ~~ classrel)  | 
| 
35201
 
c2ddb9395436
axclass: more precise treatment of naming vs. binding;
 
wenzelm 
parents: 
35021 
diff
changeset
 | 
591  | 
|> 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
 | 
592  | 
|> 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
 | 
593  | 
|> #2  | 
| 19511 | 594  | 
|> Sign.restore_naming facts_thy  | 
| 
36329
 
85004134055c
more systematic treatment of data -- avoid slightly odd nested tuples here;
 
wenzelm 
parents: 
36328 
diff
changeset
 | 
595  | 
|> map_axclasses (Symtab.update (class, axclass))  | 
| 
42383
 
0ae4ad40d7b5
simplified pretty printing context, which is only required for certain kernel operations;
 
wenzelm 
parents: 
42375 
diff
changeset
 | 
596  | 
|> map_params (fold (fn (x, _) => add_param (Context.pretty ctxt) (x, class)) params);  | 
| 19511 | 597  | 
|
598  | 
in (class, result_thy) end;  | 
|
599  | 
||
600  | 
||
| 19531 | 601  | 
|
| 19511 | 602  | 
(** axiomatizations **)  | 
603  | 
||
604  | 
local  | 
|
605  | 
||
| 36417 | 606  | 
(*old-style axioms*)  | 
| 20628 | 607  | 
fun axiomatize prep mk name add raw_args thy =  | 
608  | 
let  | 
|
609  | 
val args = prep thy raw_args;  | 
|
610  | 
val specs = mk args;  | 
|
611  | 
val names = name args;  | 
|
| 29579 | 612  | 
in  | 
613  | 
thy  | 
|
| 
42375
 
774df7c59508
report Name_Space.declare/define, relatively to context;
 
wenzelm 
parents: 
42360 
diff
changeset
 | 
614  | 
|> fold_map Thm.add_axiom_global (map Binding.name names ~~ specs)  | 
| 37249 | 615  | 
|-> fold (add o Drule.export_without_context o snd)  | 
| 29579 | 616  | 
end;  | 
| 19511 | 617  | 
|
618  | 
fun ax_classrel prep =  | 
|
| 20628 | 619  | 
axiomatize (map o prep) (map Logic.mk_classrel)  | 
| 37249 | 620  | 
(map (prefix classrel_prefix o Logic.name_classrel)) add_classrel;  | 
| 19511 | 621  | 
|
622  | 
fun ax_arity prep =  | 
|
| 42360 | 623  | 
axiomatize (prep o Proof_Context.init_global) Logic.mk_arities  | 
| 37249 | 624  | 
(map (prefix arity_prefix) o Logic.name_arities) add_arity;  | 
| 19511 | 625  | 
|
| 19706 | 626  | 
fun class_const c =  | 
627  | 
(Logic.const_of_class c, Term.itselfT (Term.aT []) --> propT);  | 
|
628  | 
||
| 19511 | 629  | 
fun ax_class prep_class prep_classrel (bclass, raw_super) thy =  | 
630  | 
let  | 
|
| 
30344
 
10a67c5ddddb
more uniform handling of binding in targets and derived elements;
 
wenzelm 
parents: 
30280 
diff
changeset
 | 
631  | 
val class = Sign.full_name thy bclass;  | 
| 24731 | 632  | 
val super = map (prep_class thy) raw_super |> Sign.minimize_sort thy;  | 
| 19511 | 633  | 
in  | 
634  | 
thy  | 
|
635  | 
|> Sign.primitive_class (bclass, super)  | 
|
636  | 
|> ax_classrel prep_classrel (map (fn c => (class, c)) super)  | 
|
| 
42375
 
774df7c59508
report Name_Space.declare/define, relatively to context;
 
wenzelm 
parents: 
42360 
diff
changeset
 | 
637  | 
|> Theory.add_deps_global "" (class_const class) (map class_const super)  | 
| 19511 | 638  | 
end;  | 
639  | 
||
640  | 
in  | 
|
641  | 
||
| 
24929
 
408becab067e
renamed AxClass.get_definition to AxClass.get_info (again);
 
wenzelm 
parents: 
24920 
diff
changeset
 | 
642  | 
val axiomatize_class = ax_class Sign.certify_class cert_classrel;  | 
| 
42375
 
774df7c59508
report Name_Space.declare/define, relatively to context;
 
wenzelm 
parents: 
42360 
diff
changeset
 | 
643  | 
val axiomatize_class_cmd =  | 
| 
 
774df7c59508
report Name_Space.declare/define, relatively to context;
 
wenzelm 
parents: 
42360 
diff
changeset
 | 
644  | 
ax_class (Proof_Context.read_class o Proof_Context.init_global) read_classrel;  | 
| 
24929
 
408becab067e
renamed AxClass.get_definition to AxClass.get_info (again);
 
wenzelm 
parents: 
24920 
diff
changeset
 | 
645  | 
val axiomatize_classrel = ax_classrel cert_classrel;  | 
| 
 
408becab067e
renamed AxClass.get_definition to AxClass.get_info (again);
 
wenzelm 
parents: 
24920 
diff
changeset
 | 
646  | 
val axiomatize_classrel_cmd = ax_classrel read_classrel;  | 
| 42360 | 647  | 
val axiomatize_arity = ax_arity Proof_Context.cert_arity;  | 
648  | 
val axiomatize_arity_cmd = ax_arity Proof_Context.read_arity;  | 
|
| 19511 | 649  | 
|
650  | 
end;  | 
|
651  | 
||
652  | 
end;  |