| author | wenzelm | 
| Tue, 16 Feb 2010 13:35:42 +0100 | |
| changeset 35144 | 8b8302da3a55 | 
| parent 33553 | 35f2b30593a8 | 
| child 35743 | c506c029a082 | 
| permissions | -rw-r--r-- | 
| 31723 
f5cafe803b55
discontinued ancient tradition to suffix certain ML module names with "_package"
 haftmann parents: 
31136diff
changeset | 1 | (* Title: HOL/Tools/typecopy.ML | 
| 20426 | 2 | Author: Florian Haftmann, TU Muenchen | 
| 3 | ||
| 20855 | 4 | Introducing copies of types using trivial typedefs; datatype-like abstraction. | 
| 20426 | 5 | *) | 
| 6 | ||
| 31723 
f5cafe803b55
discontinued ancient tradition to suffix certain ML module names with "_package"
 haftmann parents: 
31136diff
changeset | 7 | signature TYPECOPY = | 
| 20426 | 8 | sig | 
| 31735 | 9 |   type info = { vs: (string * sort) list, constr: string, typ: typ,
 | 
| 10 | inject: thm, proj: string * typ, proj_def: thm } | |
| 31136 
85d04515abb3
tuned and generalized construction of code equations for eq; tuned interface
 haftmann parents: 
31090diff
changeset | 11 | val typecopy: binding * string list -> typ -> (binding * binding) option | 
| 20483 
04aa552a83bc
TypedefPackage.add_typedef_* now yields name of introduced type constructor
 haftmann parents: 
20426diff
changeset | 12 | -> theory -> (string * info) * theory | 
| 25506 | 13 | val get_info: theory -> string -> info option | 
| 25489 | 14 | val interpretation: (string -> theory -> theory) -> theory -> theory | 
| 31136 
85d04515abb3
tuned and generalized construction of code equations for eq; tuned interface
 haftmann parents: 
31090diff
changeset | 15 | val add_default_code: string -> theory -> theory | 
| 20426 | 16 | val setup: theory -> theory | 
| 17 | end; | |
| 18 | ||
| 31723 
f5cafe803b55
discontinued ancient tradition to suffix certain ML module names with "_package"
 haftmann parents: 
31136diff
changeset | 19 | structure Typecopy: TYPECOPY = | 
| 20426 | 20 | struct | 
| 21 | ||
| 22 | (* theory data *) | |
| 23 | ||
| 24 | type info = {
 | |
| 25 | vs: (string * sort) list, | |
| 26 | constr: string, | |
| 27 | typ: typ, | |
| 28 | inject: thm, | |
| 29 | proj: string * typ, | |
| 30 | proj_def: thm | |
| 31 | }; | |
| 32 | ||
| 33522 | 33 | structure TypecopyData = Theory_Data | 
| 22846 | 34 | ( | 
| 24626 
85eceef2edc7
introduced generic concepts for theory interpretators
 haftmann parents: 
24624diff
changeset | 35 | type T = info Symtab.table; | 
| 
85eceef2edc7
introduced generic concepts for theory interpretators
 haftmann parents: 
24624diff
changeset | 36 | val empty = Symtab.empty; | 
| 20426 | 37 | val extend = I; | 
| 33522 | 38 | fun merge data = Symtab.merge (K true) data; | 
| 22846 | 39 | ); | 
| 40 | ||
| 25506 | 41 | val get_info = Symtab.lookup o TypecopyData.get; | 
| 20426 | 42 | |
| 43 | ||
| 25489 | 44 | (* interpretation of type copies *) | 
| 20426 | 45 | |
| 33314 | 46 | structure Typecopy_Interpretation = Interpretation(type T = string val eq = op =); | 
| 47 | val interpretation = Typecopy_Interpretation.interpretation; | |
| 20426 | 48 | |
| 31136 
85d04515abb3
tuned and generalized construction of code equations for eq; tuned interface
 haftmann parents: 
31090diff
changeset | 49 | |
| 31735 | 50 | (* introducing typecopies *) | 
| 31136 
85d04515abb3
tuned and generalized construction of code equations for eq; tuned interface
 haftmann parents: 
31090diff
changeset | 51 | |
| 
85d04515abb3
tuned and generalized construction of code equations for eq; tuned interface
 haftmann parents: 
31090diff
changeset | 52 | fun typecopy (raw_tyco, raw_vs) raw_ty constr_proj thy = | 
| 20426 | 53 | let | 
| 28664 | 54 | val ty = Sign.certify_typ thy raw_ty; | 
| 29272 | 55 | val vs = | 
| 56 | AList.make (the_default HOLogic.typeS o AList.lookup (op =) (Term.add_tfreesT ty [])) raw_vs; | |
| 20426 | 57 | val tac = Tactic.rtac UNIV_witness 1; | 
| 20483 
04aa552a83bc
TypedefPackage.add_typedef_* now yields name of introduced type constructor
 haftmann parents: 
20426diff
changeset | 58 |     fun add_info tyco ( { abs_type = ty_abs, rep_type = ty_rep, Abs_name = c_abs,
 | 
| 20426 | 59 | Rep_name = c_rep, Abs_inject = inject, | 
| 31723 
f5cafe803b55
discontinued ancient tradition to suffix certain ML module names with "_package"
 haftmann parents: 
31136diff
changeset | 60 | Abs_inverse = inverse, ... } : Typedef.info ) thy = | 
| 20426 | 61 | let | 
| 62 | val exists_thm = | |
| 63 | UNIV_I | |
| 64 | |> Drule.instantiate' [SOME (ctyp_of thy (Logic.varifyT ty_rep))] []; | |
| 65 | val inject' = inject OF [exists_thm, exists_thm]; | |
| 66 | val proj_def = inverse OF [exists_thm]; | |
| 67 |           val info = {
 | |
| 68 | vs = vs, | |
| 69 | constr = c_abs, | |
| 70 | typ = ty_rep, | |
| 71 | inject = inject', | |
| 72 | proj = (c_rep, ty_abs --> ty_rep), | |
| 73 | proj_def = proj_def | |
| 74 | }; | |
| 75 | in | |
| 76 | thy | |
| 24626 
85eceef2edc7
introduced generic concepts for theory interpretators
 haftmann parents: 
24624diff
changeset | 77 | |> (TypecopyData.map o Symtab.update_new) (tyco, info) | 
| 33314 | 78 | |> Typecopy_Interpretation.data tyco | 
| 20483 
04aa552a83bc
TypedefPackage.add_typedef_* now yields name of introduced type constructor
 haftmann parents: 
20426diff
changeset | 79 | |> pair (tyco, info) | 
| 20426 | 80 | end | 
| 81 | in | |
| 82 | thy | |
| 31723 
f5cafe803b55
discontinued ancient tradition to suffix certain ML module names with "_package"
 haftmann parents: 
31136diff
changeset | 83 | |> Typedef.add_typedef false (SOME raw_tyco) (raw_tyco, map fst vs, NoSyn) | 
| 26475 
3cc1e48d0ce1
eliminated quiet_mode ref of some packages (avoid CRITICAL setmp!);
 wenzelm parents: 
25864diff
changeset | 84 | (HOLogic.mk_UNIV ty) (Option.map swap constr_proj) tac | 
| 20483 
04aa552a83bc
TypedefPackage.add_typedef_* now yields name of introduced type constructor
 haftmann parents: 
20426diff
changeset | 85 | |-> (fn (tyco, info) => add_info tyco info) | 
| 20426 | 86 | end; | 
| 87 | ||
| 88 | ||
| 31136 
85d04515abb3
tuned and generalized construction of code equations for eq; tuned interface
 haftmann parents: 
31090diff
changeset | 89 | (* default code setup *) | 
| 20835 | 90 | |
| 31136 
85d04515abb3
tuned and generalized construction of code equations for eq; tuned interface
 haftmann parents: 
31090diff
changeset | 91 | fun add_default_code tyco thy = | 
| 25489 | 92 | let | 
| 32378 | 93 |     val SOME { constr = c, proj = (proj, _), proj_def = proj_eq, vs = vs,
 | 
| 31136 
85d04515abb3
tuned and generalized construction of code equations for eq; tuned interface
 haftmann parents: 
31090diff
changeset | 94 | typ = ty_rep, ... } = get_info thy tyco; | 
| 31723 
f5cafe803b55
discontinued ancient tradition to suffix certain ML module names with "_package"
 haftmann parents: 
31136diff
changeset | 95 |     val SOME { Rep_inject = proj_inject, ... } = Typedef.get_info thy tyco;
 | 
| 32378 | 96 | val constr = (c, Logic.unvarifyT (Sign.the_const_type thy c)); | 
| 28423 | 97 | val ty = Type (tyco, map TFree vs); | 
| 31136 
85d04515abb3
tuned and generalized construction of code equations for eq; tuned interface
 haftmann parents: 
31090diff
changeset | 98 | val proj = Const (proj, ty --> ty_rep); | 
| 
85d04515abb3
tuned and generalized construction of code equations for eq; tuned interface
 haftmann parents: 
31090diff
changeset | 99 |     val (t_x, t_y) = (Free ("x", ty), Free ("y", ty));
 | 
| 
85d04515abb3
tuned and generalized construction of code equations for eq; tuned interface
 haftmann parents: 
31090diff
changeset | 100 |     val eq_lhs = Const (@{const_name eq_class.eq}, ty --> ty --> HOLogic.boolT)
 | 
| 28664 | 101 | $ t_x $ t_y; | 
| 31136 
85d04515abb3
tuned and generalized construction of code equations for eq; tuned interface
 haftmann parents: 
31090diff
changeset | 102 | val eq_rhs = HOLogic.mk_eq (proj $ t_x, proj $ t_y); | 
| 
85d04515abb3
tuned and generalized construction of code equations for eq; tuned interface
 haftmann parents: 
31090diff
changeset | 103 | val eq = (HOLogic.mk_Trueprop o HOLogic.mk_eq) (eq_lhs, eq_rhs); | 
| 
85d04515abb3
tuned and generalized construction of code equations for eq; tuned interface
 haftmann parents: 
31090diff
changeset | 104 | fun tac eq_thm = Class.intro_classes_tac [] | 
| 
85d04515abb3
tuned and generalized construction of code equations for eq; tuned interface
 haftmann parents: 
31090diff
changeset | 105 | THEN (Simplifier.rewrite_goals_tac | 
| 
85d04515abb3
tuned and generalized construction of code equations for eq; tuned interface
 haftmann parents: 
31090diff
changeset | 106 |         (map Simpdata.mk_eq [eq_thm, @{thm eq}, proj_inject]))
 | 
| 
85d04515abb3
tuned and generalized construction of code equations for eq; tuned interface
 haftmann parents: 
31090diff
changeset | 107 |           THEN ALLGOALS (rtac @{thm refl});
 | 
| 28664 | 108 |     fun mk_eq_refl thy = @{thm HOL.eq_refl}
 | 
| 109 | |> Thm.instantiate | |
| 110 |          ([pairself (Thm.ctyp_of thy) (TVar (("'a", 0), @{sort eq}), Logic.varifyT ty)], [])
 | |
| 111 | |> AxClass.unoverload thy; | |
| 25489 | 112 | in | 
| 25506 | 113 | thy | 
| 28664 | 114 | |> Code.add_datatype [constr] | 
| 115 | |> Code.add_eqn proj_eq | |
| 33553 | 116 | |> Theory_Target.instantiation ([tyco], vs, [HOLogic.class_eq]) | 
| 31136 
85d04515abb3
tuned and generalized construction of code equations for eq; tuned interface
 haftmann parents: 
31090diff
changeset | 117 | |> `(fn lthy => Syntax.check_term lthy eq) | 
| 
85d04515abb3
tuned and generalized construction of code equations for eq; tuned interface
 haftmann parents: 
31090diff
changeset | 118 | |-> (fn eq => Specification.definition | 
| 
85d04515abb3
tuned and generalized construction of code equations for eq; tuned interface
 haftmann parents: 
31090diff
changeset | 119 | (NONE, (Attrib.empty_binding, eq))) | 
| 
85d04515abb3
tuned and generalized construction of code equations for eq; tuned interface
 haftmann parents: 
31090diff
changeset | 120 | |-> (fn (_, (_, eq_thm)) => | 
| 28664 | 121 | Class.prove_instantiation_exit_result Morphism.thm | 
| 31136 
85d04515abb3
tuned and generalized construction of code equations for eq; tuned interface
 haftmann parents: 
31090diff
changeset | 122 | (fn _ => fn eq_thm => tac eq_thm) eq_thm) | 
| 
85d04515abb3
tuned and generalized construction of code equations for eq; tuned interface
 haftmann parents: 
31090diff
changeset | 123 | |-> (fn eq_thm => Code.add_eqn eq_thm) | 
| 
85d04515abb3
tuned and generalized construction of code equations for eq; tuned interface
 haftmann parents: 
31090diff
changeset | 124 | |> (fn thy => Code.add_nbe_eqn (mk_eq_refl thy) thy) | 
| 25489 | 125 | end; | 
| 20426 | 126 | |
| 25489 | 127 | val setup = | 
| 33314 | 128 | Typecopy_Interpretation.init | 
| 31136 
85d04515abb3
tuned and generalized construction of code equations for eq; tuned interface
 haftmann parents: 
31090diff
changeset | 129 | #> interpretation add_default_code | 
| 20426 | 130 | |
| 22846 | 131 | end; |