| author | wenzelm | 
| Thu, 30 May 2013 16:31:53 +0200 | |
| changeset 52240 | 066c2ff17f7c | 
| parent 50768 | 2172f82de515 | 
| child 52622 | e0ff1625e96d | 
| permissions | -rw-r--r-- | 
| 37789 | 1  | 
(* Author: Alexander Krauss, TU Muenchen  | 
2  | 
Author: Christian Sternagel, University of Innsbruck  | 
|
3  | 
||
4  | 
Ad-hoc overloading of constants based on their types.  | 
|
5  | 
*)  | 
|
6  | 
||
7  | 
signature ADHOC_OVERLOADING =  | 
|
8  | 
sig  | 
|
9  | 
val add_overloaded: string -> theory -> theory  | 
|
10  | 
val add_variant: string -> string -> theory -> theory  | 
|
11  | 
||
| 45422 | 12  | 
val show_variants: bool Config.T  | 
| 37789 | 13  | 
val setup: theory -> theory  | 
14  | 
end  | 
|
15  | 
||
16  | 
structure Adhoc_Overloading: ADHOC_OVERLOADING =  | 
|
17  | 
struct  | 
|
18  | 
||
| 45422 | 19  | 
val show_variants = Attrib.setup_config_bool @{binding show_variants} (K false);
 | 
| 37789 | 20  | 
|
21  | 
||
22  | 
(* errors *)  | 
|
23  | 
||
24  | 
fun duplicate_variant_err int_name ext_name =  | 
|
25  | 
  error ("Constant " ^ quote int_name ^ " is already a variant of " ^ quote ext_name);
 | 
|
26  | 
||
27  | 
fun not_overloaded_err name =  | 
|
28  | 
  error ("Constant " ^ quote name ^ " is not declared as overloaded");
 | 
|
29  | 
||
30  | 
fun already_overloaded_err name =  | 
|
31  | 
  error ("Constant " ^ quote name ^ " is already declared as overloaded");
 | 
|
32  | 
||
33  | 
fun unresolved_err ctxt (c, T) t reason =  | 
|
34  | 
  error ("Unresolved overloading of  " ^ quote c ^ " :: " ^
 | 
|
35  | 
quote (Syntax.string_of_typ ctxt T) ^ " in " ^  | 
|
36  | 
    quote (Syntax.string_of_term ctxt t) ^ " (" ^ reason ^ ")");
 | 
|
37  | 
||
38  | 
||
39  | 
(* theory data *)  | 
|
40  | 
||
41  | 
structure Overload_Data = Theory_Data  | 
|
42  | 
(  | 
|
43  | 
type T =  | 
|
44  | 
    { internalize : (string * typ) list Symtab.table,
 | 
|
45  | 
externalize : string Symtab.table };  | 
|
46  | 
  val empty = {internalize=Symtab.empty, externalize=Symtab.empty};
 | 
|
47  | 
val extend = I;  | 
|
48  | 
||
49  | 
fun merge_ext int_name (ext_name1, ext_name2) =  | 
|
50  | 
if ext_name1 = ext_name2 then ext_name1  | 
|
51  | 
else duplicate_variant_err int_name ext_name1;  | 
|
52  | 
||
| 
41472
 
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
 
wenzelm 
parents: 
37818 
diff
changeset
 | 
53  | 
fun merge  | 
| 
 
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
 
wenzelm 
parents: 
37818 
diff
changeset
 | 
54  | 
    ({internalize = int1, externalize = ext1},
 | 
| 
 
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
 
wenzelm 
parents: 
37818 
diff
changeset
 | 
55  | 
      {internalize = int2, externalize = ext2}) : T =
 | 
| 
50768
 
2172f82de515
tuned -- prefer high-level Table.merge with its slightly more conservative update;
 
wenzelm 
parents: 
45444 
diff
changeset
 | 
56  | 
    {internalize = Symtab.merge_list (op =) (int1, int2),
 | 
| 
41472
 
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
 
wenzelm 
parents: 
37818 
diff
changeset
 | 
57  | 
externalize = Symtab.join merge_ext (ext1, ext2)};  | 
| 37789 | 58  | 
);  | 
59  | 
||
60  | 
fun map_tables f g =  | 
|
61  | 
  Overload_Data.map (fn {internalize=int, externalize=ext} =>
 | 
|
62  | 
    {internalize=f int, externalize=g ext});
 | 
|
63  | 
||
64  | 
val is_overloaded = Symtab.defined o #internalize o Overload_Data.get;  | 
|
65  | 
val get_variants = Symtab.lookup o #internalize o Overload_Data.get;  | 
|
66  | 
val get_external = Symtab.lookup o #externalize o Overload_Data.get;  | 
|
67  | 
||
68  | 
fun add_overloaded ext_name thy =  | 
|
69  | 
let val _ = not (is_overloaded thy ext_name) orelse already_overloaded_err ext_name;  | 
|
70  | 
in map_tables (Symtab.update (ext_name, [])) I thy end;  | 
|
71  | 
||
72  | 
fun add_variant ext_name name thy =  | 
|
73  | 
let  | 
|
74  | 
val _ = is_overloaded thy ext_name orelse not_overloaded_err ext_name;  | 
|
| 45422 | 75  | 
val _ =  | 
76  | 
(case get_external thy name of  | 
|
77  | 
NONE => ()  | 
|
78  | 
| SOME gen' => duplicate_variant_err name gen');  | 
|
| 37789 | 79  | 
val T = Sign.the_const_type thy name;  | 
80  | 
in  | 
|
81  | 
map_tables (Symtab.cons_list (ext_name, (name, T)))  | 
|
| 45422 | 82  | 
(Symtab.update (name, ext_name)) thy  | 
| 37789 | 83  | 
end  | 
84  | 
||
85  | 
||
86  | 
(* check / uncheck *)  | 
|
87  | 
||
88  | 
fun unifiable_with ctxt T1 (c, T2) =  | 
|
89  | 
let  | 
|
| 42361 | 90  | 
val thy = Proof_Context.theory_of ctxt;  | 
| 37789 | 91  | 
val maxidx1 = Term.maxidx_of_typ T1;  | 
92  | 
val T2' = Logic.incr_tvar (maxidx1 + 1) T2;  | 
|
93  | 
val maxidx2 = Int.max (maxidx1, Term.maxidx_of_typ T2');  | 
|
94  | 
in  | 
|
95  | 
(Sign.typ_unify thy (T1, T2') (Vartab.empty, maxidx2); SOME c)  | 
|
96  | 
handle Type.TUNIFY => NONE  | 
|
97  | 
end;  | 
|
98  | 
||
99  | 
fun insert_internal_same ctxt t (Const (c, T)) =  | 
|
| 45422 | 100  | 
(case map_filter (unifiable_with ctxt T)  | 
101  | 
(Same.function (get_variants (Proof_Context.theory_of ctxt)) c) of  | 
|
102  | 
[] => unresolved_err ctxt (c, T) t "no instances"  | 
|
103  | 
| [c'] => Const (c', dummyT)  | 
|
104  | 
| _ => raise Same.SAME)  | 
|
| 37789 | 105  | 
| insert_internal_same _ _ _ = raise Same.SAME;  | 
106  | 
||
107  | 
fun insert_external_same ctxt _ (Const (c, T)) =  | 
|
| 45422 | 108  | 
Const (Same.function (get_external (Proof_Context.theory_of ctxt)) c, T)  | 
| 37789 | 109  | 
| insert_external_same _ _ _ = raise Same.SAME;  | 
110  | 
||
111  | 
fun gen_check_uncheck replace ts ctxt =  | 
|
112  | 
Same.capture (Same.map (fn t => Term_Subst.map_aterms_same (replace ctxt t) t)) ts  | 
|
113  | 
|> Option.map (rpair ctxt);  | 
|
114  | 
||
115  | 
val check = gen_check_uncheck insert_internal_same;  | 
|
| 45422 | 116  | 
|
117  | 
fun uncheck ts ctxt =  | 
|
118  | 
if Config.get ctxt show_variants then NONE  | 
|
| 37789 | 119  | 
else gen_check_uncheck insert_external_same ts ctxt;  | 
120  | 
||
121  | 
fun reject_unresolved ts ctxt =  | 
|
122  | 
let  | 
|
| 42361 | 123  | 
val thy = Proof_Context.theory_of ctxt;  | 
| 37789 | 124  | 
fun check_unresolved t =  | 
| 45422 | 125  | 
(case filter (is_overloaded thy o fst) (Term.add_consts t []) of  | 
126  | 
[] => ()  | 
|
127  | 
| ((c, T) :: _) => unresolved_err ctxt (c, T) t "multiple instances");  | 
|
| 37789 | 128  | 
val _ = map check_unresolved ts;  | 
129  | 
in NONE end;  | 
|
130  | 
||
131  | 
||
132  | 
(* setup *)  | 
|
133  | 
||
| 45422 | 134  | 
val setup = Context.theory_map  | 
| 45444 | 135  | 
(Syntax_Phases.term_check' 0 "adhoc_overloading" check  | 
136  | 
#> Syntax_Phases.term_check' 1 "adhoc_overloading_unresolved_check" reject_unresolved  | 
|
137  | 
#> Syntax_Phases.term_uncheck' 0 "adhoc_overloading" uncheck);  | 
|
| 37789 | 138  | 
|
| 45422 | 139  | 
end;  |