| author | aspinall | 
| Wed, 22 Nov 2006 12:01:59 +0100 | |
| changeset 21464 | abaf43b011ee | 
| parent 21463 | 42dd50268c8b | 
| child 21687 | f689f729afab | 
| permissions | -rw-r--r-- | 
| 
404
 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 
wenzelm 
parents:  
diff
changeset
 | 
1  | 
(* Title: Pure/axclass.ML  | 
| 
 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 
wenzelm 
parents:  
diff
changeset
 | 
2  | 
ID: $Id$  | 
| 
 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 
wenzelm 
parents:  
diff
changeset
 | 
3  | 
Author: Markus Wenzel, TU Muenchen  | 
| 
 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 
wenzelm 
parents:  
diff
changeset
 | 
4  | 
|
| 19511 | 5  | 
Type classes as parameter records and predicates, with explicit  | 
6  | 
definitions and proofs.  | 
|
| 
404
 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 
wenzelm 
parents:  
diff
changeset
 | 
7  | 
*)  | 
| 
 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 
wenzelm 
parents:  
diff
changeset
 | 
8  | 
|
| 
 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 
wenzelm 
parents:  
diff
changeset
 | 
9  | 
signature AX_CLASS =  | 
| 3938 | 10  | 
sig  | 
| 
21463
 
42dd50268c8b
completed class parameter handling in axclass.ML
 
haftmann 
parents: 
20628 
diff
changeset
 | 
11  | 
  val get_definition: theory -> class -> {def: thm, intro: thm, axioms: thm list,
 | 
| 
 
42dd50268c8b
completed class parameter handling in axclass.ML
 
haftmann 
parents: 
20628 
diff
changeset
 | 
12  | 
params: (string * typ) list, operational: bool}  | 
| 
19243
 
5dcb899a8486
moved read_class, read/cert_classrel/arity to sign.ML;
 
wenzelm 
parents: 
19134 
diff
changeset
 | 
13  | 
val class_intros: theory -> thm list  | 
| 20107 | 14  | 
val class_of_param: theory -> string -> class option  | 
| 
21463
 
42dd50268c8b
completed class parameter handling in axclass.ML
 
haftmann 
parents: 
20628 
diff
changeset
 | 
15  | 
val params_of_class: theory -> class -> string * (string * typ) list  | 
| 19511 | 16  | 
val print_axclasses: theory -> unit  | 
| 19405 | 17  | 
val cert_classrel: theory -> class * class -> class * class  | 
18  | 
val read_classrel: theory -> xstring * xstring -> class * class  | 
|
19  | 
val add_classrel: thm -> theory -> theory  | 
|
20  | 
val add_arity: thm -> theory -> theory  | 
|
21  | 
val prove_classrel: class * class -> tactic -> theory -> theory  | 
|
| 
19243
 
5dcb899a8486
moved read_class, read/cert_classrel/arity to sign.ML;
 
wenzelm 
parents: 
19134 
diff
changeset
 | 
22  | 
val prove_arity: string * sort list * sort -> tactic -> theory -> theory  | 
| 19511 | 23  | 
val define_class: bstring * xstring list -> string list ->  | 
| 
19418
 
03b01c9314fc
reworded add_axclass(_i): canonical specification format,
 
wenzelm 
parents: 
19405 
diff
changeset
 | 
24  | 
((bstring * Attrib.src list) * string list) list -> theory -> class * theory  | 
| 19511 | 25  | 
val define_class_i: bstring * class list -> string list ->  | 
| 
19418
 
03b01c9314fc
reworded add_axclass(_i): canonical specification format,
 
wenzelm 
parents: 
19405 
diff
changeset
 | 
26  | 
((bstring * attribute list) * term list) list -> theory -> class * theory  | 
| 19511 | 27  | 
val axiomatize_class: bstring * xstring list -> theory -> theory  | 
28  | 
val axiomatize_class_i: bstring * class list -> theory -> theory  | 
|
29  | 
val axiomatize_classrel: (xstring * xstring) list -> theory -> theory  | 
|
30  | 
val axiomatize_classrel_i: (class * class) list -> theory -> theory  | 
|
31  | 
val axiomatize_arity: xstring * string list * string -> theory -> theory  | 
|
32  | 
val axiomatize_arity_i: arity -> theory -> theory  | 
|
| 19574 | 33  | 
type cache  | 
34  | 
val cache: cache  | 
|
35  | 
val of_sort: theory -> typ * sort -> cache -> thm list * cache (*exception Sorts.CLASS_ERROR*)  | 
|
| 3938 | 36  | 
end;  | 
| 
404
 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 
wenzelm 
parents:  
diff
changeset
 | 
37  | 
|
| 15801 | 38  | 
structure AxClass: AX_CLASS =  | 
| 
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  | 
|
| 19405 | 43  | 
(* class parameters (canonical order) *)  | 
| 423 | 44  | 
|
| 19405 | 45  | 
type param = string * class;  | 
| 423 | 46  | 
|
| 19405 | 47  | 
fun add_param pp ((x, c): param) params =  | 
48  | 
(case AList.lookup (op =) params x of  | 
|
49  | 
NONE => (x, c) :: params  | 
|
50  | 
  | SOME c' => error ("Duplicate class parameter " ^ quote x ^
 | 
|
51  | 
" for " ^ Pretty.string_of_sort pp [c] ^  | 
|
52  | 
(if c = c' then "" else " and " ^ Pretty.string_of_sort pp [c'])));  | 
|
| 423 | 53  | 
|
| 19405 | 54  | 
fun merge_params _ ([], qs) = qs  | 
55  | 
| merge_params pp (ps, qs) =  | 
|
56  | 
fold_rev (fn q => if member (op =) ps q then I else add_param pp q) qs ps;  | 
|
| 423 | 57  | 
|
58  | 
||
| 19511 | 59  | 
(* axclasses *)  | 
| 6379 | 60  | 
|
| 
19243
 
5dcb899a8486
moved read_class, read/cert_classrel/arity to sign.ML;
 
wenzelm 
parents: 
19134 
diff
changeset
 | 
61  | 
val introN = "intro";  | 
| 19511 | 62  | 
val superN = "super";  | 
| 
19243
 
5dcb899a8486
moved read_class, read/cert_classrel/arity to sign.ML;
 
wenzelm 
parents: 
19134 
diff
changeset
 | 
63  | 
val axiomsN = "axioms";  | 
| 
 
5dcb899a8486
moved read_class, read/cert_classrel/arity to sign.ML;
 
wenzelm 
parents: 
19134 
diff
changeset
 | 
64  | 
|
| 
21463
 
42dd50268c8b
completed class parameter handling in axclass.ML
 
haftmann 
parents: 
20628 
diff
changeset
 | 
65  | 
val param_tyvarname = "'a";  | 
| 
 
42dd50268c8b
completed class parameter handling in axclass.ML
 
haftmann 
parents: 
20628 
diff
changeset
 | 
66  | 
|
| 19392 | 67  | 
datatype axclass = AxClass of  | 
| 19460 | 68  | 
 {def: thm,
 | 
| 19392 | 69  | 
intro: thm,  | 
| 
21463
 
42dd50268c8b
completed class parameter handling in axclass.ML
 
haftmann 
parents: 
20628 
diff
changeset
 | 
70  | 
axioms: thm list,  | 
| 
 
42dd50268c8b
completed class parameter handling in axclass.ML
 
haftmann 
parents: 
20628 
diff
changeset
 | 
71  | 
params: (string * typ) list,  | 
| 
 
42dd50268c8b
completed class parameter handling in axclass.ML
 
haftmann 
parents: 
20628 
diff
changeset
 | 
72  | 
operational: bool (* == at least one class operation,  | 
| 
 
42dd50268c8b
completed class parameter handling in axclass.ML
 
haftmann 
parents: 
20628 
diff
changeset
 | 
73  | 
or at least two operational superclasses *)};  | 
| 19392 | 74  | 
|
| 19460 | 75  | 
type axclasses = axclass Symtab.table * param list;  | 
| 19392 | 76  | 
|
| 
21463
 
42dd50268c8b
completed class parameter handling in axclass.ML
 
haftmann 
parents: 
20628 
diff
changeset
 | 
77  | 
fun make_axclass ((def, intro, axioms), (params, operational)) =  | 
| 
 
42dd50268c8b
completed class parameter handling in axclass.ML
 
haftmann 
parents: 
20628 
diff
changeset
 | 
78  | 
  AxClass {def = def, intro = intro, axioms = axioms, params = params,
 | 
| 
 
42dd50268c8b
completed class parameter handling in axclass.ML
 
haftmann 
parents: 
20628 
diff
changeset
 | 
79  | 
operational = operational};  | 
| 19405 | 80  | 
|
81  | 
fun merge_axclasses pp ((tab1, params1), (tab2, params2)) : axclasses =  | 
|
82  | 
(Symtab.merge (K true) (tab1, tab2), merge_params pp (params1, params2));  | 
|
83  | 
||
| 19392 | 84  | 
|
85  | 
(* instances *)  | 
|
86  | 
||
| 20628 | 87  | 
val classrel_prefix = "classrel_";  | 
88  | 
val arity_prefix = "arity_";  | 
|
| 19511 | 89  | 
|
| 19574 | 90  | 
type instances =  | 
91  | 
((class * class) * thm) list *  | 
|
92  | 
((class * sort list) * thm) list Symtab.table;  | 
|
| 6379 | 93  | 
|
| 19574 | 94  | 
fun merge_instances ((classrel1, arities1): instances, (classrel2, arities2)) =  | 
95  | 
(merge (eq_fst op =) (classrel1, classrel2),  | 
|
96  | 
Symtab.join (K (merge (eq_fst op =))) (arities1, arities2));  | 
|
| 19392 | 97  | 
|
98  | 
||
| 19511 | 99  | 
(* setup data *)  | 
| 19392 | 100  | 
|
101  | 
structure AxClassData = TheoryDataFun  | 
|
| 16458 | 102  | 
(struct  | 
| 19392 | 103  | 
val name = "Pure/axclass";  | 
| 19574 | 104  | 
type T = axclasses * instances;  | 
105  | 
val empty : T = ((Symtab.empty, []), ([], Symtab.empty));  | 
|
106  | 
val copy = I;  | 
|
107  | 
val extend = I;  | 
|
108  | 
fun merge pp ((axclasses1, instances1), (axclasses2, instances2)) =  | 
|
109  | 
(merge_axclasses pp (axclasses1, axclasses2), (merge_instances (instances1, instances2)));  | 
|
| 19460 | 110  | 
fun print _ _ = ();  | 
| 16458 | 111  | 
end);  | 
| 6379 | 112  | 
|
| 19392 | 113  | 
val _ = Context.add_setup AxClassData.init;  | 
| 6379 | 114  | 
|
115  | 
||
| 19574 | 116  | 
(* maintain axclasses *)  | 
| 19392 | 117  | 
|
| 19574 | 118  | 
val get_axclasses = #1 o AxClassData.get;  | 
119  | 
fun map_axclasses f = AxClassData.map (apfst f);  | 
|
120  | 
||
121  | 
val lookup_def = Symtab.lookup o #1 o get_axclasses;  | 
|
| 6379 | 122  | 
|
| 19511 | 123  | 
fun get_definition thy c =  | 
124  | 
(case lookup_def thy c of  | 
|
| 19392 | 125  | 
SOME (AxClass info) => info  | 
| 19511 | 126  | 
  | NONE => error ("Undefined type class " ^ quote c));
 | 
| 6379 | 127  | 
|
| 19123 | 128  | 
fun class_intros thy =  | 
| 19392 | 129  | 
let  | 
130  | 
fun add_intro c =  | 
|
| 19511 | 131  | 
      (case lookup_def thy c of SOME (AxClass {intro, ...}) => cons intro | _ => I);
 | 
| 19392 | 132  | 
val classes = Sign.classes thy;  | 
133  | 
in map (Thm.class_triv thy) classes @ fold add_intro classes [] end;  | 
|
| 
15876
 
a67343c6ab2a
sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
 
wenzelm 
parents: 
15853 
diff
changeset
 | 
134  | 
|
| 
 
a67343c6ab2a
sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
 
wenzelm 
parents: 
15853 
diff
changeset
 | 
135  | 
|
| 19460 | 136  | 
fun get_params thy pred =  | 
| 19574 | 137  | 
let val params = #2 (get_axclasses thy);  | 
| 19460 | 138  | 
in fold (fn (x, c) => if pred c then cons x else I) params [] end;  | 
139  | 
||
140  | 
fun params_of thy c = get_params thy (fn c' => c' = c);  | 
|
141  | 
fun all_params_of thy S = get_params thy (fn c => Sign.subsort thy (S, [c]));  | 
|
142  | 
||
| 20107 | 143  | 
fun class_of_param thy =  | 
| 19955 | 144  | 
AList.lookup (op =) (#2 (get_axclasses thy));  | 
145  | 
||
| 
21463
 
42dd50268c8b
completed class parameter handling in axclass.ML
 
haftmann 
parents: 
20628 
diff
changeset
 | 
146  | 
fun params_of_class thy class =  | 
| 
 
42dd50268c8b
completed class parameter handling in axclass.ML
 
haftmann 
parents: 
20628 
diff
changeset
 | 
147  | 
(param_tyvarname, (#params o get_definition thy) class);  | 
| 19460 | 148  | 
|
| 19511 | 149  | 
(* maintain instances *)  | 
| 19503 | 150  | 
|
| 19574 | 151  | 
val get_instances = #2 o AxClassData.get;  | 
152  | 
fun map_instances f = AxClassData.map (apsnd f);  | 
|
| 19528 | 153  | 
|
154  | 
||
| 19574 | 155  | 
fun the_classrel thy (c1, c2) =  | 
156  | 
(case AList.lookup (op =) (#1 (get_instances thy)) (c1, c2) of  | 
|
157  | 
SOME th => Thm.transfer thy th  | 
|
158  | 
  | NONE => error ("Unproven class relation " ^ Sign.string_of_classrel thy [c1, c2]));
 | 
|
159  | 
||
160  | 
fun put_classrel arg = map_instances (fn (classrel, arities) =>  | 
|
161  | 
(insert (eq_fst op =) arg classrel, arities));  | 
|
| 19503 | 162  | 
|
163  | 
||
| 19574 | 164  | 
fun the_arity thy a (c, Ss) =  | 
165  | 
(case AList.lookup (op =) (Symtab.lookup_list (#2 (get_instances thy)) a) (c, Ss) of  | 
|
166  | 
SOME th => Thm.transfer thy th  | 
|
167  | 
  | NONE => error ("Unproven type arity " ^ Sign.string_of_arity thy (a, Ss, [c])));
 | 
|
| 19503 | 168  | 
|
| 19574 | 169  | 
fun put_arity ((t, Ss, c), th) = map_instances (fn (classrel, arities) =>  | 
170  | 
(classrel, arities |> Symtab.insert_list (eq_fst op =) (t, ((c, Ss), th))));  | 
|
| 19503 | 171  | 
|
172  | 
||
| 19511 | 173  | 
(* print data *)  | 
| 19460 | 174  | 
|
175  | 
fun print_axclasses thy =  | 
|
176  | 
let  | 
|
| 19574 | 177  | 
val axclasses = #1 (get_axclasses thy);  | 
| 19460 | 178  | 
val ctxt = ProofContext.init thy;  | 
179  | 
||
| 
21463
 
42dd50268c8b
completed class parameter handling in axclass.ML
 
haftmann 
parents: 
20628 
diff
changeset
 | 
180  | 
    fun pretty_axclass (class, AxClass {def, intro, axioms, params, operational}) =
 | 
| 19460 | 181  | 
Pretty.block (Pretty.fbreaks  | 
182  | 
[Pretty.block  | 
|
| 19511 | 183  | 
[Pretty.str "class ", ProofContext.pretty_sort ctxt [class], Pretty.str ":"],  | 
| 19460 | 184  | 
        Pretty.strs ("parameters:" :: params_of thy class),
 | 
185  | 
        ProofContext.pretty_fact ctxt ("def", [def]),
 | 
|
186  | 
ProofContext.pretty_fact ctxt (introN, [intro]),  | 
|
187  | 
ProofContext.pretty_fact ctxt (axiomsN, axioms)]);  | 
|
188  | 
in Pretty.writeln (Pretty.chunks (map pretty_axclass (Symtab.dest axclasses))) end;  | 
|
189  | 
||
190  | 
||
| 
404
 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 
wenzelm 
parents:  
diff
changeset
 | 
191  | 
|
| 19511 | 192  | 
(** instances **)  | 
| 
19418
 
03b01c9314fc
reworded add_axclass(_i): canonical specification format,
 
wenzelm 
parents: 
19405 
diff
changeset
 | 
193  | 
|
| 
 
03b01c9314fc
reworded add_axclass(_i): canonical specification format,
 
wenzelm 
parents: 
19405 
diff
changeset
 | 
194  | 
(* class relations *)  | 
| 
 
03b01c9314fc
reworded add_axclass(_i): canonical specification format,
 
wenzelm 
parents: 
19405 
diff
changeset
 | 
195  | 
|
| 19405 | 196  | 
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
 | 
197  | 
let  | 
| 19405 | 198  | 
val (c1, c2) = pairself (Sign.certify_class thy) raw_rel;  | 
| 19511 | 199  | 
val _ = Type.add_classrel (Sign.pp thy) (c1, c2) (Sign.tsig_of thy);  | 
| 19405 | 200  | 
val _ =  | 
| 19460 | 201  | 
(case subtract (op =) (all_params_of thy [c1]) (all_params_of thy [c2]) of  | 
| 19405 | 202  | 
[] => ()  | 
203  | 
      | xs => raise TYPE ("Class " ^ Sign.string_of_sort thy [c1] ^ " lacks parameter(s) " ^
 | 
|
204  | 
commas_quote xs ^ " of " ^ Sign.string_of_sort thy [c2], [], []));  | 
|
205  | 
in (c1, c2) end;  | 
|
206  | 
||
207  | 
fun read_classrel thy raw_rel =  | 
|
208  | 
cert_classrel thy (pairself (Sign.read_class thy) raw_rel)  | 
|
209  | 
handle TYPE (msg, _, _) => error msg;  | 
|
210  | 
||
211  | 
||
212  | 
(* primitive rules *)  | 
|
213  | 
||
214  | 
fun add_classrel th thy =  | 
|
215  | 
let  | 
|
216  | 
    fun err () = raise THM ("add_classrel: malformed class relation", 0, [th]);
 | 
|
217  | 
val prop = Drule.plain_prop_of (Thm.transfer thy th);  | 
|
218  | 
val rel = Logic.dest_classrel prop handle TERM _ => err ();  | 
|
219  | 
val (c1, c2) = cert_classrel thy rel handle TYPE _ => err ();  | 
|
| 19574 | 220  | 
in  | 
221  | 
thy  | 
|
222  | 
|> Sign.primitive_classrel (c1, c2)  | 
|
223  | 
|> put_classrel ((c1, c2), Drule.unconstrainTs th)  | 
|
224  | 
end;  | 
|
| 
15876
 
a67343c6ab2a
sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
 
wenzelm 
parents: 
15853 
diff
changeset
 | 
225  | 
|
| 19405 | 226  | 
fun add_arity th thy =  | 
| 
15876
 
a67343c6ab2a
sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
 
wenzelm 
parents: 
15853 
diff
changeset
 | 
227  | 
let  | 
| 19528 | 228  | 
    fun err () = raise THM ("add_arity: malformed type arity", 0, [th]);
 | 
| 19405 | 229  | 
val prop = Drule.plain_prop_of (Thm.transfer thy th);  | 
| 19528 | 230  | 
val (t, Ss, c) = Logic.dest_arity prop handle TERM _ => err ();  | 
231  | 
val _ = if map (Sign.certify_sort thy) Ss = Ss then () else err ();  | 
|
| 19574 | 232  | 
in  | 
233  | 
thy  | 
|
234  | 
|> Sign.primitive_arity (t, Ss, [c])  | 
|
235  | 
|> put_arity ((t, Ss, c), Drule.unconstrainTs th)  | 
|
236  | 
end;  | 
|
| 
404
 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 
wenzelm 
parents:  
diff
changeset
 | 
237  | 
|
| 
 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 
wenzelm 
parents:  
diff
changeset
 | 
238  | 
|
| 
19243
 
5dcb899a8486
moved read_class, read/cert_classrel/arity to sign.ML;
 
wenzelm 
parents: 
19134 
diff
changeset
 | 
239  | 
(* tactical proofs *)  | 
| 
15876
 
a67343c6ab2a
sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
 
wenzelm 
parents: 
15853 
diff
changeset
 | 
240  | 
|
| 19405 | 241  | 
fun prove_classrel raw_rel tac thy =  | 
| 
15876
 
a67343c6ab2a
sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
 
wenzelm 
parents: 
15853 
diff
changeset
 | 
242  | 
let  | 
| 19405 | 243  | 
val (c1, c2) = cert_classrel thy raw_rel;  | 
| 20049 | 244  | 
val th = Goal.prove (ProofContext.init thy) [] []  | 
245  | 
(Logic.mk_classrel (c1, c2)) (fn _ => tac) handle ERROR msg =>  | 
|
| 
19243
 
5dcb899a8486
moved read_class, read/cert_classrel/arity to sign.ML;
 
wenzelm 
parents: 
19134 
diff
changeset
 | 
246  | 
      cat_error msg ("The error(s) above occurred while trying to prove class relation " ^
 | 
| 
 
5dcb899a8486
moved read_class, read/cert_classrel/arity to sign.ML;
 
wenzelm 
parents: 
19134 
diff
changeset
 | 
247  | 
quote (Sign.string_of_classrel thy [c1, c2]));  | 
| 19511 | 248  | 
in  | 
249  | 
thy  | 
|
| 20628 | 250  | 
|> PureThy.add_thms [((prefix classrel_prefix (Logic.name_classrel (c1, c2)), th), [])]  | 
| 19511 | 251  | 
|-> (fn [th'] => add_classrel th')  | 
252  | 
end;  | 
|
| 
404
 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 
wenzelm 
parents:  
diff
changeset
 | 
253  | 
|
| 
19243
 
5dcb899a8486
moved read_class, read/cert_classrel/arity to sign.ML;
 
wenzelm 
parents: 
19134 
diff
changeset
 | 
254  | 
fun prove_arity raw_arity tac thy =  | 
| 
15876
 
a67343c6ab2a
sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
 
wenzelm 
parents: 
15853 
diff
changeset
 | 
255  | 
let  | 
| 
19243
 
5dcb899a8486
moved read_class, read/cert_classrel/arity to sign.ML;
 
wenzelm 
parents: 
19134 
diff
changeset
 | 
256  | 
val arity = Sign.cert_arity thy raw_arity;  | 
| 20628 | 257  | 
val names = map (prefix arity_prefix) (Logic.name_arities arity);  | 
| 19405 | 258  | 
val props = Logic.mk_arities arity;  | 
| 20049 | 259  | 
val ths = Goal.prove_multi (ProofContext.init thy) [] [] props  | 
| 18678 | 260  | 
(fn _ => Tactic.precise_conjunction_tac (length props) 1 THEN tac) handle ERROR msg =>  | 
| 
19243
 
5dcb899a8486
moved read_class, read/cert_classrel/arity to sign.ML;
 
wenzelm 
parents: 
19134 
diff
changeset
 | 
261  | 
        cat_error msg ("The error(s) above occurred while trying to prove type arity " ^
 | 
| 
 
5dcb899a8486
moved read_class, read/cert_classrel/arity to sign.ML;
 
wenzelm 
parents: 
19134 
diff
changeset
 | 
262  | 
quote (Sign.string_of_arity thy arity));  | 
| 19511 | 263  | 
in  | 
264  | 
thy  | 
|
| 20628 | 265  | 
|> PureThy.add_thms (map (rpair []) (names ~~ ths))  | 
| 19511 | 266  | 
|-> fold add_arity  | 
267  | 
end;  | 
|
268  | 
||
269  | 
||
270  | 
||
271  | 
(** class definitions **)  | 
|
272  | 
||
273  | 
local  | 
|
274  | 
||
| 
21463
 
42dd50268c8b
completed class parameter handling in axclass.ML
 
haftmann 
parents: 
20628 
diff
changeset
 | 
275  | 
fun read_param thy raw_t =  | 
| 
 
42dd50268c8b
completed class parameter handling in axclass.ML
 
haftmann 
parents: 
20628 
diff
changeset
 | 
276  | 
let  | 
| 
 
42dd50268c8b
completed class parameter handling in axclass.ML
 
haftmann 
parents: 
20628 
diff
changeset
 | 
277  | 
val t = Sign.read_term thy raw_t  | 
| 
 
42dd50268c8b
completed class parameter handling in axclass.ML
 
haftmann 
parents: 
20628 
diff
changeset
 | 
278  | 
in case try dest_Const t  | 
| 
 
42dd50268c8b
completed class parameter handling in axclass.ML
 
haftmann 
parents: 
20628 
diff
changeset
 | 
279  | 
of SOME (c, _) => c  | 
| 
 
42dd50268c8b
completed class parameter handling in axclass.ML
 
haftmann 
parents: 
20628 
diff
changeset
 | 
280  | 
    | NONE => error ("Not a constant: " ^ Sign.string_of_term thy t)
 | 
| 
 
42dd50268c8b
completed class parameter handling in axclass.ML
 
haftmann 
parents: 
20628 
diff
changeset
 | 
281  | 
end;  | 
| 
 
42dd50268c8b
completed class parameter handling in axclass.ML
 
haftmann 
parents: 
20628 
diff
changeset
 | 
282  | 
|
| 
 
42dd50268c8b
completed class parameter handling in axclass.ML
 
haftmann 
parents: 
20628 
diff
changeset
 | 
283  | 
fun def_class prep_class prep_att prep_param prep_propp  | 
| 
 
42dd50268c8b
completed class parameter handling in axclass.ML
 
haftmann 
parents: 
20628 
diff
changeset
 | 
284  | 
(bclass, raw_super) raw_params raw_specs thy =  | 
| 19511 | 285  | 
let  | 
286  | 
val ctxt = ProofContext.init thy;  | 
|
287  | 
val pp = ProofContext.pp ctxt;  | 
|
288  | 
||
289  | 
||
290  | 
(* prepare specification *)  | 
|
291  | 
||
292  | 
val bconst = Logic.const_of_class bclass;  | 
|
293  | 
val class = Sign.full_name thy bclass;  | 
|
294  | 
val super = map (prep_class thy) raw_super |> Sign.certify_sort thy;  | 
|
295  | 
||
296  | 
fun prep_axiom t =  | 
|
297  | 
(case Term.add_tfrees t [] of  | 
|
298  | 
[(a, S)] =>  | 
|
299  | 
if Sign.subsort thy (super, S) then t  | 
|
300  | 
          else error ("Sort constraint of type variable " ^
 | 
|
301  | 
setmp show_sorts true (Pretty.string_of_typ pp) (TFree (a, S)) ^  | 
|
302  | 
" needs to be weaker than " ^ Pretty.string_of_sort pp super)  | 
|
303  | 
| [] => t  | 
|
304  | 
      | _ => error ("Multiple type variables in class axiom:\n" ^ Pretty.string_of_term pp t))
 | 
|
| 
20548
 
8ef25fe585a8
renamed Term.map_term_types to Term.map_types (cf. Term.fold_types);
 
wenzelm 
parents: 
20260 
diff
changeset
 | 
305  | 
|> map_types (Term.map_atyps (fn TFree _ => Term.aT [] | U => U))  | 
| 19511 | 306  | 
|> Logic.close_form;  | 
307  | 
||
| 19585 | 308  | 
val axiomss = prep_propp (ctxt, map (map (rpair []) o snd) raw_specs)  | 
| 19511 | 309  | 
|> snd |> map (map (prep_axiom o fst));  | 
310  | 
val name_atts = Attrib.map_specs (prep_att thy) raw_specs |> map fst;  | 
|
311  | 
||
312  | 
||
313  | 
(* definition *)  | 
|
314  | 
||
315  | 
val conjs = map (curry Logic.mk_inclass (Term.aT [])) super @ flat axiomss;  | 
|
316  | 
val class_eq =  | 
|
317  | 
Logic.mk_equals (Logic.mk_inclass (Term.aT [], class), Logic.mk_conjunction_list conjs);  | 
|
318  | 
||
319  | 
val ([def], def_thy) =  | 
|
320  | 
thy  | 
|
321  | 
|> Sign.primitive_class (bclass, super)  | 
|
322  | 
|> PureThy.add_defs_i false [((Thm.def_name bconst, class_eq), [])];  | 
|
323  | 
val (raw_intro, (raw_classrel, raw_axioms)) =  | 
|
324  | 
(Conjunction.split_defined (length conjs) def) ||> chop (length super);  | 
|
| 19392 | 325  | 
|
| 
19418
 
03b01c9314fc
reworded add_axclass(_i): canonical specification format,
 
wenzelm 
parents: 
19405 
diff
changeset
 | 
326  | 
|
| 19511 | 327  | 
(* facts *)  | 
328  | 
||
329  | 
val class_triv = Thm.class_triv def_thy class;  | 
|
330  | 
val ([(_, [intro]), (_, classrel), (_, axioms)], facts_thy) =  | 
|
331  | 
def_thy  | 
|
332  | 
|> PureThy.note_thmss_qualified "" bconst  | 
|
333  | 
[((introN, []), [([Drule.standard raw_intro], [])]),  | 
|
334  | 
((superN, []), [(map Drule.standard raw_classrel, [])]),  | 
|
335  | 
((axiomsN, []), [(map (fn th => Drule.standard (class_triv RS th)) raw_axioms, [])])];  | 
|
336  | 
||
| 
21463
 
42dd50268c8b
completed class parameter handling in axclass.ML
 
haftmann 
parents: 
20628 
diff
changeset
 | 
337  | 
(* params *)  | 
| 
 
42dd50268c8b
completed class parameter handling in axclass.ML
 
haftmann 
parents: 
20628 
diff
changeset
 | 
338  | 
|
| 
 
42dd50268c8b
completed class parameter handling in axclass.ML
 
haftmann 
parents: 
20628 
diff
changeset
 | 
339  | 
val params = map (prep_param thy) raw_params;  | 
| 
 
42dd50268c8b
completed class parameter handling in axclass.ML
 
haftmann 
parents: 
20628 
diff
changeset
 | 
340  | 
val params_typs = map (fn param =>  | 
| 
 
42dd50268c8b
completed class parameter handling in axclass.ML
 
haftmann 
parents: 
20628 
diff
changeset
 | 
341  | 
let  | 
| 
 
42dd50268c8b
completed class parameter handling in axclass.ML
 
haftmann 
parents: 
20628 
diff
changeset
 | 
342  | 
val ty = Sign.the_const_type thy param;  | 
| 
 
42dd50268c8b
completed class parameter handling in axclass.ML
 
haftmann 
parents: 
20628 
diff
changeset
 | 
343  | 
val var = case Term.typ_tvars ty  | 
| 
 
42dd50268c8b
completed class parameter handling in axclass.ML
 
haftmann 
parents: 
20628 
diff
changeset
 | 
344  | 
of [(v, _)] => v  | 
| 
 
42dd50268c8b
completed class parameter handling in axclass.ML
 
haftmann 
parents: 
20628 
diff
changeset
 | 
345  | 
          | _ => error ("exactly one type variable required in parameter " ^ quote param);
 | 
| 
 
42dd50268c8b
completed class parameter handling in axclass.ML
 
haftmann 
parents: 
20628 
diff
changeset
 | 
346  | 
val ty' = Term.typ_subst_TVars [(var, TFree (param_tyvarname, []))] ty;  | 
| 
 
42dd50268c8b
completed class parameter handling in axclass.ML
 
haftmann 
parents: 
20628 
diff
changeset
 | 
347  | 
in (param, ty') end) params;  | 
| 
 
42dd50268c8b
completed class parameter handling in axclass.ML
 
haftmann 
parents: 
20628 
diff
changeset
 | 
348  | 
val operational = length params_typs > 0 orelse  | 
| 
 
42dd50268c8b
completed class parameter handling in axclass.ML
 
haftmann 
parents: 
20628 
diff
changeset
 | 
349  | 
length (filter (the_default false o Option.map  | 
| 
 
42dd50268c8b
completed class parameter handling in axclass.ML
 
haftmann 
parents: 
20628 
diff
changeset
 | 
350  | 
        (fn AxClass { operational, ... } => operational) o lookup_def thy) super) > 1;
 | 
| 19511 | 351  | 
|
352  | 
(* result *)  | 
|
353  | 
||
354  | 
val result_thy =  | 
|
355  | 
facts_thy  | 
|
| 19574 | 356  | 
|> fold put_classrel (map (pair class) super ~~ classrel)  | 
| 19511 | 357  | 
|> Sign.add_path bconst  | 
358  | 
|> PureThy.note_thmss_i "" (name_atts ~~ map Thm.simple_fact (unflat axiomss axioms)) |> snd  | 
|
359  | 
|> Sign.restore_naming facts_thy  | 
|
| 19574 | 360  | 
|> map_axclasses (fn (axclasses, parameters) =>  | 
| 
21463
 
42dd50268c8b
completed class parameter handling in axclass.ML
 
haftmann 
parents: 
20628 
diff
changeset
 | 
361  | 
(Symtab.update (class, make_axclass ((def, intro, axioms), (params_typs, operational))) axclasses,  | 
| 19574 | 362  | 
fold (fn x => add_param pp (x, class)) params parameters));  | 
| 19511 | 363  | 
|
364  | 
in (class, result_thy) end;  | 
|
365  | 
||
366  | 
in  | 
|
367  | 
||
| 
21463
 
42dd50268c8b
completed class parameter handling in axclass.ML
 
haftmann 
parents: 
20628 
diff
changeset
 | 
368  | 
val define_class = def_class Sign.read_class Attrib.attribute read_param ProofContext.read_propp;  | 
| 
 
42dd50268c8b
completed class parameter handling in axclass.ML
 
haftmann 
parents: 
20628 
diff
changeset
 | 
369  | 
val define_class_i = def_class Sign.certify_class (K I) (K I) ProofContext.cert_propp;  | 
| 19511 | 370  | 
|
371  | 
end;  | 
|
372  | 
||
373  | 
||
| 19531 | 374  | 
|
| 19511 | 375  | 
(** axiomatizations **)  | 
376  | 
||
377  | 
local  | 
|
378  | 
||
| 20628 | 379  | 
fun axiomatize prep mk name add raw_args thy =  | 
380  | 
let  | 
|
381  | 
val args = prep thy raw_args;  | 
|
382  | 
val specs = mk args;  | 
|
383  | 
val names = name args;  | 
|
384  | 
in thy |> PureThy.add_axioms_i (map (rpair []) (names ~~ specs)) |-> fold add end;  | 
|
| 19511 | 385  | 
|
386  | 
fun ax_classrel prep =  | 
|
| 20628 | 387  | 
axiomatize (map o prep) (map Logic.mk_classrel)  | 
388  | 
(map (prefix classrel_prefix o Logic.name_classrel)) add_classrel;  | 
|
| 19511 | 389  | 
|
390  | 
fun ax_arity prep =  | 
|
| 20628 | 391  | 
axiomatize prep Logic.mk_arities  | 
392  | 
(map (prefix arity_prefix) o Logic.name_arities) add_arity;  | 
|
| 19511 | 393  | 
|
| 19706 | 394  | 
fun class_const c =  | 
395  | 
(Logic.const_of_class c, Term.itselfT (Term.aT []) --> propT);  | 
|
396  | 
||
| 19511 | 397  | 
fun ax_class prep_class prep_classrel (bclass, raw_super) thy =  | 
398  | 
let  | 
|
399  | 
val class = Sign.full_name thy bclass;  | 
|
400  | 
val super = map (prep_class thy) raw_super |> Sign.certify_sort thy;  | 
|
401  | 
in  | 
|
402  | 
thy  | 
|
403  | 
|> Sign.primitive_class (bclass, super)  | 
|
404  | 
|> ax_classrel prep_classrel (map (fn c => (class, c)) super)  | 
|
| 19706 | 405  | 
|> Theory.add_deps "" (class_const class) (map class_const super)  | 
| 19511 | 406  | 
end;  | 
407  | 
||
408  | 
in  | 
|
409  | 
||
410  | 
val axiomatize_class = ax_class Sign.read_class read_classrel;  | 
|
411  | 
val axiomatize_class_i = ax_class Sign.certify_class cert_classrel;  | 
|
412  | 
val axiomatize_classrel = ax_classrel read_classrel;  | 
|
413  | 
val axiomatize_classrel_i = ax_classrel cert_classrel;  | 
|
414  | 
val axiomatize_arity = ax_arity Sign.read_arity;  | 
|
415  | 
val axiomatize_arity_i = ax_arity Sign.cert_arity;  | 
|
416  | 
||
417  | 
end;  | 
|
418  | 
||
419  | 
||
420  | 
||
421  | 
(** explicit derivations -- cached **)  | 
|
422  | 
||
| 19574 | 423  | 
datatype cache = Types of (class * thm) list Typtab.table;  | 
424  | 
val cache = Types Typtab.empty;  | 
|
425  | 
||
| 19511 | 426  | 
local  | 
| 19503 | 427  | 
|
| 19574 | 428  | 
fun lookup_type (Types cache) = AList.lookup (op =) o Typtab.lookup_list cache;  | 
429  | 
fun insert_type T der (Types cache) = Types (Typtab.insert_list (eq_fst op =) (T, der) cache);  | 
|
| 19503 | 430  | 
|
| 19522 | 431  | 
fun derive_type _ (_, []) = []  | 
432  | 
| derive_type thy (typ, sort) =  | 
|
433  | 
let  | 
|
| 19528 | 434  | 
val vars = Term.fold_atyps  | 
| 19522 | 435  | 
(fn T as TFree (_, S) => insert (eq_fst op =) (T, S)  | 
436  | 
| T as TVar (_, S) => insert (eq_fst op =) (T, S)  | 
|
| 19528 | 437  | 
| _ => I) typ [];  | 
| 19574 | 438  | 
val hyps = vars  | 
439  | 
|> map (fn (T, S) => (T, Drule.sort_triv thy (T, S) ~~ S));  | 
|
| 19642 | 440  | 
val ths = (typ, sort) |> Sorts.of_sort_derivation (Sign.pp thy) (Sign.classes_of thy)  | 
| 19574 | 441  | 
           {classrel =
 | 
442  | 
fn (th, c1) => fn c2 => th RS the_classrel thy (c1, c2),  | 
|
443  | 
constructor =  | 
|
444  | 
fn a => fn dom => fn c =>  | 
|
445  | 
let val Ss = map (map snd) dom and ths = maps (map fst) dom  | 
|
446  | 
in ths MRS the_arity thy a (c, Ss) end,  | 
|
447  | 
variable =  | 
|
448  | 
the_default [] o AList.lookup (op =) hyps};  | 
|
449  | 
in ths end;  | 
|
| 19503 | 450  | 
|
| 19511 | 451  | 
in  | 
452  | 
||
| 19574 | 453  | 
fun of_sort thy (typ, sort) cache =  | 
454  | 
let  | 
|
455  | 
val sort' = filter (is_none o lookup_type cache typ) sort;  | 
|
456  | 
val ths' = derive_type thy (typ, sort')  | 
|
457  | 
      handle ERROR msg => cat_error msg ("The error(s) above occurred for sort derivation: " ^
 | 
|
458  | 
Sign.string_of_typ thy typ ^ " :: " ^ Sign.string_of_sort thy sort');  | 
|
459  | 
val cache' = cache |> fold (insert_type typ) (sort' ~~ ths');  | 
|
460  | 
val ths =  | 
|
461  | 
sort |> map (fn c =>  | 
|
462  | 
Goal.finish (the (lookup_type cache' typ c) RS  | 
|
463  | 
Goal.init (Thm.cterm_of thy (Logic.mk_inclass (typ, c))))  | 
|
| 20260 | 464  | 
|> Thm.adjust_maxidx_thm ~1);  | 
| 19574 | 465  | 
in (ths, cache') end;  | 
| 19503 | 466  | 
|
| 19511 | 467  | 
end;  | 
| 
19418
 
03b01c9314fc
reworded add_axclass(_i): canonical specification format,
 
wenzelm 
parents: 
19405 
diff
changeset
 | 
468  | 
|
| 
15876
 
a67343c6ab2a
sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
 
wenzelm 
parents: 
15853 
diff
changeset
 | 
469  | 
end;  |