| author | haftmann | 
| Fri, 20 Jul 2007 14:28:05 +0200 | |
| changeset 23880 | 64b9806e160b | 
| parent 22846 | fb79144af9a3 | 
| child 24219 | e558fe311376 | 
| permissions | -rw-r--r-- | 
| 20426 | 1 | (* Title: HOL/Tools/typecopy_package.ML | 
| 2 | ID: $Id$ | |
| 3 | Author: Florian Haftmann, TU Muenchen | |
| 4 | ||
| 20855 | 5 | Introducing copies of types using trivial typedefs; datatype-like abstraction. | 
| 20426 | 6 | *) | 
| 7 | ||
| 8 | signature TYPECOPY_PACKAGE = | |
| 9 | sig | |
| 10 |   type info = {
 | |
| 11 | vs: (string * sort) list, | |
| 12 | constr: string, | |
| 13 | typ: typ, | |
| 14 | inject: thm, | |
| 15 | proj: string * typ, | |
| 16 | proj_def: thm | |
| 17 | } | |
| 18 | val add_typecopy: bstring * string list -> typ -> (bstring * bstring) option | |
| 20483 
04aa552a83bc
TypedefPackage.add_typedef_* now yields name of introduced type constructor
 haftmann parents: 
20426diff
changeset | 19 | -> theory -> (string * info) * theory | 
| 20426 | 20 | val get_typecopies: theory -> string list | 
| 21 | val get_typecopy_info: theory -> string -> info option | |
| 20835 | 22 | type hook = string * info -> theory -> theory | 
| 23 | val add_hook: hook -> theory -> theory | |
| 24 | val get_spec: theory -> string -> (string * sort) list * (string * typ list) list | |
| 25 | val get_eq: theory -> string -> thm | |
| 22846 | 26 | val print_typecopies: theory -> unit | 
| 20426 | 27 | val setup: theory -> theory | 
| 28 | end; | |
| 29 | ||
| 30 | structure TypecopyPackage: TYPECOPY_PACKAGE = | |
| 31 | struct | |
| 32 | ||
| 33 | (* theory data *) | |
| 34 | ||
| 35 | type info = {
 | |
| 36 | vs: (string * sort) list, | |
| 37 | constr: string, | |
| 38 | typ: typ, | |
| 39 | inject: thm, | |
| 40 | proj: string * typ, | |
| 41 | proj_def: thm | |
| 42 | }; | |
| 43 | ||
| 44 | type hook = string * info -> theory -> theory; | |
| 45 | ||
| 46 | structure TypecopyData = TheoryDataFun | |
| 22846 | 47 | ( | 
| 20426 | 48 | type T = info Symtab.table * (serial * hook) list; | 
| 49 | val empty = (Symtab.empty, []); | |
| 50 | val copy = I; | |
| 51 | val extend = I; | |
| 52 | fun merge _ ((tab1, hooks1), (tab2, hooks2) : T) = | |
| 53 | (Symtab.merge (K true) (tab1, tab2), AList.merge (op =) (K true) (hooks1, hooks2)); | |
| 22846 | 54 | ); | 
| 55 | ||
| 56 | fun print_typecopies thy = | |
| 57 | let | |
| 58 | val (tab, _) = TypecopyData.get thy; | |
| 59 |     fun mk (tyco, { vs, constr, typ, proj = (proj, _), ... } : info) =
 | |
| 60 | (Pretty.block o Pretty.breaks) [ | |
| 61 | Sign.pretty_typ thy (Type (tyco, map TFree vs)), | |
| 62 | Pretty.str "=", | |
| 63 | (Pretty.str o Sign.extern_const thy) constr, | |
| 64 | Sign.pretty_typ thy typ, | |
| 65 |         Pretty.block [Pretty.str "(", (Pretty.str o Sign.extern_const thy) proj, Pretty.str  ")"]];
 | |
| 20426 | 66 | in | 
| 22846 | 67 | (Pretty.writeln o Pretty.block o Pretty.fbreaks) | 
| 68 | (Pretty.str "type copies:" :: map mk (Symtab.dest tab)) | |
| 20426 | 69 | end; | 
| 70 | ||
| 71 | val get_typecopies = Symtab.keys o fst o TypecopyData.get; | |
| 72 | val get_typecopy_info = Symtab.lookup o fst o TypecopyData.get; | |
| 73 | ||
| 74 | ||
| 75 | (* hook management *) | |
| 76 | ||
| 77 | fun add_hook hook = | |
| 78 | (TypecopyData.map o apsnd o cons) (serial (), hook); | |
| 79 | ||
| 80 | fun invoke_hooks tyco_info thy = | |
| 81 | fold (fn (_, f) => f tyco_info) ((snd o TypecopyData.get) thy) thy; | |
| 82 | ||
| 83 | ||
| 84 | (* add a type copy *) | |
| 85 | ||
| 86 | local | |
| 87 | ||
| 20483 
04aa552a83bc
TypedefPackage.add_typedef_* now yields name of introduced type constructor
 haftmann parents: 
20426diff
changeset | 88 | fun gen_add_typecopy prep_typ (raw_tyco, raw_vs) raw_ty constr_proj thy = | 
| 20426 | 89 | let | 
| 90 | val ty = prep_typ thy raw_ty; | |
| 91 | val vs = AList.make (the_default HOLogic.typeS o AList.lookup (op =) (typ_tfrees ty)) raw_vs; | |
| 92 | val tac = Tactic.rtac UNIV_witness 1; | |
| 20483 
04aa552a83bc
TypedefPackage.add_typedef_* now yields name of introduced type constructor
 haftmann parents: 
20426diff
changeset | 93 |     fun add_info tyco ( { abs_type = ty_abs, rep_type = ty_rep, Abs_name = c_abs,
 | 
| 20426 | 94 | Rep_name = c_rep, Abs_inject = inject, | 
| 21708 | 95 | Abs_inverse = inverse, ... } : TypedefPackage.info ) thy = | 
| 20426 | 96 | let | 
| 97 | val exists_thm = | |
| 98 | UNIV_I | |
| 99 | |> Drule.instantiate' [SOME (ctyp_of thy (Logic.varifyT ty_rep))] []; | |
| 100 | val inject' = inject OF [exists_thm, exists_thm]; | |
| 101 | val proj_def = inverse OF [exists_thm]; | |
| 102 |           val info = {
 | |
| 103 | vs = vs, | |
| 104 | constr = c_abs, | |
| 105 | typ = ty_rep, | |
| 106 | inject = inject', | |
| 107 | proj = (c_rep, ty_abs --> ty_rep), | |
| 108 | proj_def = proj_def | |
| 109 | }; | |
| 110 | in | |
| 111 | thy | |
| 20483 
04aa552a83bc
TypedefPackage.add_typedef_* now yields name of introduced type constructor
 haftmann parents: 
20426diff
changeset | 112 | |> (TypecopyData.map o apfst o Symtab.update_new) (tyco, info) | 
| 
04aa552a83bc
TypedefPackage.add_typedef_* now yields name of introduced type constructor
 haftmann parents: 
20426diff
changeset | 113 | |> invoke_hooks (tyco, info) | 
| 
04aa552a83bc
TypedefPackage.add_typedef_* now yields name of introduced type constructor
 haftmann parents: 
20426diff
changeset | 114 | |> pair (tyco, info) | 
| 20426 | 115 | end | 
| 116 | in | |
| 117 | thy | |
| 20483 
04aa552a83bc
TypedefPackage.add_typedef_* now yields name of introduced type constructor
 haftmann parents: 
20426diff
changeset | 118 | |> setmp TypedefPackage.quiet_mode true | 
| 
04aa552a83bc
TypedefPackage.add_typedef_* now yields name of introduced type constructor
 haftmann parents: 
20426diff
changeset | 119 | (TypedefPackage.add_typedef_i false (SOME raw_tyco) (raw_tyco, map fst vs, NoSyn) | 
| 
04aa552a83bc
TypedefPackage.add_typedef_* now yields name of introduced type constructor
 haftmann parents: 
20426diff
changeset | 120 | (HOLogic.mk_UNIV ty) (Option.map swap constr_proj)) tac | 
| 
04aa552a83bc
TypedefPackage.add_typedef_* now yields name of introduced type constructor
 haftmann parents: 
20426diff
changeset | 121 | |-> (fn (tyco, info) => add_info tyco info) | 
| 20426 | 122 | end; | 
| 123 | ||
| 124 | in | |
| 125 | ||
| 126 | val add_typecopy = gen_add_typecopy Sign.certify_typ; | |
| 127 | ||
| 22846 | 128 | end; | 
| 20426 | 129 | |
| 130 | ||
| 22423 | 131 | (* equality function equation and datatype specification *) | 
| 20835 | 132 | |
| 133 | fun get_eq thy tyco = | |
| 134 | (#inject o the o get_typecopy_info thy) tyco; | |
| 135 | ||
| 136 | fun get_spec thy tyco = | |
| 137 | let | |
| 138 |     val SOME { vs, constr, typ, ... } = get_typecopy_info thy tyco
 | |
| 139 | in (vs, [(constr, [typ])]) end; | |
| 140 | ||
| 141 | ||
| 142 | (* hook for projection function code *) | |
| 20426 | 143 | |
| 21708 | 144 | fun add_project (_ , {proj_def, ...} : info) =
 | 
| 22484 
25dfebd7b4c8
improved treatment of defining equations stemming from specification tools
 haftmann parents: 
22423diff
changeset | 145 | CodegenData.add_func true proj_def; | 
| 20426 | 146 | |
| 22846 | 147 | val setup = add_hook add_project; | 
| 20426 | 148 | |
| 22846 | 149 | end; |