| author | huffman | 
| Fri, 02 Sep 2011 20:58:31 -0700 | |
| changeset 44678 | 21eb31192850 | 
| parent 42402 | c7139609b67d | 
| child 45422 | 711dac69111b | 
| 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 | ||
| 10 | val add_overloaded: string -> theory -> theory | |
| 11 | val add_variant: string -> string -> theory -> theory | |
| 12 | ||
| 13 | val show_variants: bool Unsynchronized.ref | |
| 14 | val setup: theory -> theory | |
| 15 | ||
| 16 | end | |
| 17 | ||
| 18 | ||
| 19 | structure Adhoc_Overloading: ADHOC_OVERLOADING = | |
| 20 | struct | |
| 21 | ||
| 22 | val show_variants = Unsynchronized.ref false; | |
| 23 | ||
| 24 | ||
| 25 | (* errors *) | |
| 26 | ||
| 27 | fun duplicate_variant_err int_name ext_name = | |
| 28 |   error ("Constant " ^ quote int_name ^ " is already a variant of " ^ quote ext_name);
 | |
| 29 | ||
| 30 | fun not_overloaded_err name = | |
| 31 |   error ("Constant " ^ quote name ^ " is not declared as overloaded");
 | |
| 32 | ||
| 33 | fun already_overloaded_err name = | |
| 34 |   error ("Constant " ^ quote name ^ " is already declared as overloaded");
 | |
| 35 | ||
| 36 | fun unresolved_err ctxt (c, T) t reason = | |
| 37 |   error ("Unresolved overloading of  " ^ quote c ^ " :: " ^
 | |
| 38 | quote (Syntax.string_of_typ ctxt T) ^ " in " ^ | |
| 39 |     quote (Syntax.string_of_term ctxt t) ^ " (" ^ reason ^ ")");
 | |
| 40 | ||
| 41 | ||
| 42 | (* theory data *) | |
| 43 | ||
| 44 | structure Overload_Data = Theory_Data | |
| 45 | ( | |
| 46 | type T = | |
| 47 |     { internalize : (string * typ) list Symtab.table,
 | |
| 48 | externalize : string Symtab.table }; | |
| 49 |   val empty = {internalize=Symtab.empty, externalize=Symtab.empty};
 | |
| 50 | val extend = I; | |
| 51 | ||
| 52 | fun merge_ext int_name (ext_name1, ext_name2) = | |
| 53 | if ext_name1 = ext_name2 then ext_name1 | |
| 54 | else duplicate_variant_err int_name ext_name1; | |
| 55 | ||
| 41472 
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
 wenzelm parents: 
37818diff
changeset | 56 | fun merge | 
| 
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
 wenzelm parents: 
37818diff
changeset | 57 |     ({internalize = int1, externalize = ext1},
 | 
| 
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
 wenzelm parents: 
37818diff
changeset | 58 |       {internalize = int2, externalize = ext2}) : T =
 | 
| 
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
 wenzelm parents: 
37818diff
changeset | 59 |     {internalize = Symtab.join (K (Library.merge (op =))) (int1, int2),
 | 
| 
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
 wenzelm parents: 
37818diff
changeset | 60 | externalize = Symtab.join merge_ext (ext1, ext2)}; | 
| 37789 | 61 | ); | 
| 62 | ||
| 63 | fun map_tables f g = | |
| 64 |   Overload_Data.map (fn {internalize=int, externalize=ext} =>
 | |
| 65 |     {internalize=f int, externalize=g ext});
 | |
| 66 | ||
| 67 | val is_overloaded = Symtab.defined o #internalize o Overload_Data.get; | |
| 68 | val get_variants = Symtab.lookup o #internalize o Overload_Data.get; | |
| 69 | val get_external = Symtab.lookup o #externalize o Overload_Data.get; | |
| 70 | ||
| 71 | fun add_overloaded ext_name thy = | |
| 72 | let val _ = not (is_overloaded thy ext_name) orelse already_overloaded_err ext_name; | |
| 73 | in map_tables (Symtab.update (ext_name, [])) I thy end; | |
| 74 | ||
| 75 | fun add_variant ext_name name thy = | |
| 76 | let | |
| 77 | val _ = is_overloaded thy ext_name orelse not_overloaded_err ext_name; | |
| 78 | val _ = case get_external thy name of | |
| 79 | NONE => () | |
| 80 | | SOME gen' => duplicate_variant_err name gen'; | |
| 81 | val T = Sign.the_const_type thy name; | |
| 82 | in | |
| 83 | map_tables (Symtab.cons_list (ext_name, (name, T))) | |
| 84 | (Symtab.update (name, ext_name)) thy | |
| 85 | end | |
| 86 | ||
| 87 | ||
| 88 | (* check / uncheck *) | |
| 89 | ||
| 90 | fun unifiable_with ctxt T1 (c, T2) = | |
| 91 | let | |
| 42361 | 92 | val thy = Proof_Context.theory_of ctxt; | 
| 37789 | 93 | val maxidx1 = Term.maxidx_of_typ T1; | 
| 94 | val T2' = Logic.incr_tvar (maxidx1 + 1) T2; | |
| 95 | val maxidx2 = Int.max (maxidx1, Term.maxidx_of_typ T2'); | |
| 96 | in | |
| 97 | (Sign.typ_unify thy (T1, T2') (Vartab.empty, maxidx2); SOME c) | |
| 98 | handle Type.TUNIFY => NONE | |
| 99 | end; | |
| 100 | ||
| 101 | fun insert_internal_same ctxt t (Const (c, T)) = | |
| 102 | (case map_filter (unifiable_with ctxt T) | |
| 42361 | 103 | (Same.function (get_variants (Proof_Context.theory_of ctxt)) c) of | 
| 37789 | 104 | [] => unresolved_err ctxt (c, T) t "no instances" | 
| 105 | | [c'] => Const (c', dummyT) | |
| 106 | | _ => raise Same.SAME) | |
| 107 | | insert_internal_same _ _ _ = raise Same.SAME; | |
| 108 | ||
| 109 | fun insert_external_same ctxt _ (Const (c, T)) = | |
| 42361 | 110 | Const (Same.function (get_external (Proof_Context.theory_of ctxt)) c, T) | 
| 37789 | 111 | | insert_external_same _ _ _ = raise Same.SAME; | 
| 112 | ||
| 113 | fun gen_check_uncheck replace ts ctxt = | |
| 114 | Same.capture (Same.map (fn t => Term_Subst.map_aterms_same (replace ctxt t) t)) ts | |
| 115 | |> Option.map (rpair ctxt); | |
| 116 | ||
| 117 | val check = gen_check_uncheck insert_internal_same; | |
| 118 | fun uncheck ts ctxt = | |
| 119 | if !show_variants then NONE | |
| 120 | else gen_check_uncheck insert_external_same ts ctxt; | |
| 121 | ||
| 122 | fun reject_unresolved ts ctxt = | |
| 123 | let | |
| 42361 | 124 | val thy = Proof_Context.theory_of ctxt; | 
| 37789 | 125 | fun check_unresolved t = | 
| 126 | case filter (is_overloaded thy o fst) (Term.add_consts t []) of | |
| 127 | [] => () | |
| 128 | | ((c, T) :: _) => unresolved_err ctxt (c, T) t "multiple instances"; | |
| 129 | ||
| 130 | val _ = map check_unresolved ts; | |
| 131 | in NONE end; | |
| 132 | ||
| 133 | ||
| 134 | (* setup *) | |
| 135 | ||
| 136 | val setup = Context.theory_map | |
| 42402 
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
 wenzelm parents: 
42361diff
changeset | 137 | (Syntax.context_term_check 0 "adhoc_overloading" check | 
| 
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
 wenzelm parents: 
42361diff
changeset | 138 | #> Syntax.context_term_check 1 "adhoc_overloading_unresolved_check" reject_unresolved | 
| 
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
 wenzelm parents: 
42361diff
changeset | 139 | #> Syntax.context_term_uncheck 0 "adhoc_overloading" uncheck); | 
| 37789 | 140 | |
| 141 | end |