| author | wenzelm | 
| Tue, 17 Sep 2024 11:14:25 +0200 | |
| changeset 80894 | 3e0ca6af6738 | 
| parent 80694 | 58a209c8d40a | 
| child 81947 | 5be5b2114ecd | 
| permissions | -rw-r--r-- | 
| 31723 
f5cafe803b55
discontinued ancient tradition to suffix certain ML module names with "_package"
 haftmann parents: 
30345diff
changeset | 1 | (* Title: HOL/Tools/typedef.ML | 
| 16458 | 2 | Author: Markus Wenzel and Stefan Berghofer, TU Muenchen | 
| 4866 | 3 | |
| 21352 | 4 | Gordon/HOL-style type definitions: create a new syntactic type | 
| 35741 | 5 | represented by a non-empty set. | 
| 4866 | 6 | *) | 
| 7 | ||
| 31723 
f5cafe803b55
discontinued ancient tradition to suffix certain ML module names with "_package"
 haftmann parents: 
30345diff
changeset | 8 | signature TYPEDEF = | 
| 4866 | 9 | sig | 
| 19705 | 10 | type info = | 
| 36107 | 11 |    {rep_type: typ, abs_type: typ, Rep_name: string, Abs_name: string, axiom_name: string} *
 | 
| 49833 | 12 |    {inhabited: thm, type_definition: thm, Rep: thm, Rep_inverse: thm, Abs_inverse: thm,
 | 
| 13 | Rep_inject: thm, Abs_inject: thm, Rep_cases: thm, Abs_cases: thm, | |
| 29061 
c67cc9402ba9
inhabitance goal is now stated in original form and result contracted --
 wenzelm parents: 
29059diff
changeset | 14 | Rep_induct: thm, Abs_induct: thm} | 
| 35741 | 15 | val transform_info: morphism -> info -> info | 
| 16 | val get_info: Proof.context -> string -> info list | |
| 17 | val get_info_global: theory -> string -> info list | |
| 58662 | 18 | val interpretation: (string -> local_theory -> local_theory) -> theory -> theory | 
| 61110 | 19 |   type bindings = {Rep_name: binding, Abs_name: binding, type_definition_name: binding}
 | 
| 20 | val default_bindings: binding -> bindings | |
| 21 | val make_bindings: binding -> bindings option -> bindings | |
| 22 | val make_morphisms: binding -> (binding * binding) option -> bindings | |
| 61260 | 23 | val overloaded: bool Config.T | 
| 24 |   val add_typedef: {overloaded: bool} -> binding * (string * sort) list * mixfix ->
 | |
| 61110 | 25 | term -> bindings option -> (Proof.context -> tactic) -> local_theory -> | 
| 58959 | 26 | (string * info) * local_theory | 
| 61260 | 27 |   val add_typedef_global: {overloaded: bool} -> binding * (string * sort) list * mixfix ->
 | 
| 61110 | 28 | term -> bindings option -> (Proof.context -> tactic) -> theory -> | 
| 58959 | 29 | (string * info) * theory | 
| 61260 | 30 |   val typedef: {overloaded: bool} -> binding * (string * sort) list * mixfix ->
 | 
| 31 | term -> bindings option -> local_theory -> Proof.state | |
| 32 |   val typedef_cmd: {overloaded: bool} -> binding * (string * string option) list * mixfix ->
 | |
| 33 | string -> bindings option -> local_theory -> Proof.state | |
| 4866 | 34 | end; | 
| 35 | ||
| 31723 
f5cafe803b55
discontinued ancient tradition to suffix certain ML module names with "_package"
 haftmann parents: 
30345diff
changeset | 36 | structure Typedef: TYPEDEF = | 
| 4866 | 37 | struct | 
| 38 | ||
| 17922 | 39 | (** type definitions **) | 
| 40 | ||
| 41 | (* theory data *) | |
| 15259 | 42 | |
| 19705 | 43 | type info = | 
| 35994 
9cc3df9a606e
Typedef.info: separate global and local part, only the latter is transformed by morphisms;
 wenzelm parents: 
35840diff
changeset | 44 | (*global part*) | 
| 36107 | 45 |   {rep_type: typ, abs_type: typ, Rep_name: string, Abs_name: string, axiom_name: string} *
 | 
| 35741 | 46 | (*local part*) | 
| 49833 | 47 |   {inhabited: thm, type_definition: thm, Rep: thm, Rep_inverse: thm, Abs_inverse: thm,
 | 
| 48 | Rep_inject: thm, Abs_inject: thm, Rep_cases: thm, Abs_cases: thm, | |
| 35994 
9cc3df9a606e
Typedef.info: separate global and local part, only the latter is transformed by morphisms;
 wenzelm parents: 
35840diff
changeset | 49 | Rep_induct: thm, Abs_induct: thm}; | 
| 19459 | 50 | |
| 35741 | 51 | fun transform_info phi (info: info) = | 
| 52 | let | |
| 53 | val thm = Morphism.thm phi; | |
| 49833 | 54 |     val (global_info, {inhabited, type_definition, Rep, Rep_inverse, Abs_inverse,
 | 
| 55 | Rep_inject, Abs_inject, Rep_cases, Abs_cases, Rep_induct, Abs_induct}) = info; | |
| 35741 | 56 | in | 
| 35994 
9cc3df9a606e
Typedef.info: separate global and local part, only the latter is transformed by morphisms;
 wenzelm parents: 
35840diff
changeset | 57 | (global_info, | 
| 
9cc3df9a606e
Typedef.info: separate global and local part, only the latter is transformed by morphisms;
 wenzelm parents: 
35840diff
changeset | 58 |      {inhabited = thm inhabited, type_definition = thm type_definition,
 | 
| 49833 | 59 | Rep = thm Rep, Rep_inverse = thm Rep_inverse, Abs_inverse = thm Abs_inverse, | 
| 60 | Rep_inject = thm Rep_inject, Abs_inject = thm Abs_inject, | |
| 61 | Rep_cases = thm Rep_cases, Abs_cases = thm Abs_cases, | |
| 62 | Rep_induct = thm Rep_induct, Abs_induct = thm Abs_induct}) | |
| 35741 | 63 | end; | 
| 64 | ||
| 65 | structure Data = Generic_Data | |
| 22846 | 66 | ( | 
| 35741 | 67 | type T = info list Symtab.table; | 
| 15259 | 68 | val empty = Symtab.empty; | 
| 35741 | 69 | fun merge data = Symtab.merge_list (K true) data; | 
| 22846 | 70 | ); | 
| 15259 | 71 | |
| 61103 | 72 | fun get_info_generic context = | 
| 73 | Symtab.lookup_list (Data.get context) #> | |
| 67664 | 74 | map (transform_info (Morphism.transfer_morphism'' context)); | 
| 35741 | 75 | |
| 61103 | 76 | val get_info = get_info_generic o Context.Proof; | 
| 77 | val get_info_global = get_info_generic o Context.Theory; | |
| 78 | ||
| 79 | fun put_info name info = | |
| 80 | Data.map (Symtab.cons_list (name, transform_info Morphism.trim_context_morphism info)); | |
| 35741 | 81 | |
| 82 | ||
| 83 | (* global interpretation *) | |
| 84 | ||
| 58662 | 85 | structure Typedef_Plugin = Plugin(type T = string); | 
| 86 | ||
| 67149 | 87 | val typedef_plugin = Plugin_Name.declare_setup \<^binding>\<open>typedef\<close>; | 
| 56375 
32e0da92c786
use same idiom as used for datatype 'size' function to name constants and theorems emerging from various type interpretations -- reduces the chances of name clashes on theory merges
 blanchet parents: 
54883diff
changeset | 88 | |
| 58662 | 89 | fun interpretation f = | 
| 90 | Typedef_Plugin.interpretation typedef_plugin | |
| 91 | (fn name => fn lthy => | |
| 92 | lthy | |
| 59880 
30687c3f2b10
clarified role of naming for background theory: transform_binding (e.g. for "concealed" flag) uses naming of hypothetical context;
 wenzelm parents: 
59859diff
changeset | 93 | |> Local_Theory.map_background_naming | 
| 58662 | 94 | (Name_Space.root_path #> Name_Space.add_path (Long_Name.qualifier name)) | 
| 95 | |> f name | |
| 59880 
30687c3f2b10
clarified role of naming for background theory: transform_binding (e.g. for "concealed" flag) uses naming of hypothetical context;
 wenzelm parents: 
59859diff
changeset | 96 | |> Local_Theory.restore_background_naming lthy); | 
| 35741 | 97 | |
| 98 | ||
| 99 | (* primitive typedef axiomatization -- for fresh typedecl *) | |
| 100 | ||
| 67149 | 101 | val typedef_overloaded = Attrib.setup_config_bool \<^binding>\<open>typedef_overloaded\<close> (K false); | 
| 61260 | 102 | |
| 80631 | 103 | fun mk_inhabited T A = | 
| 80694 | 104 | \<^instantiate>\<open>'a = T and A in prop \<open>\<exists>x::'a. x \<in> A\<close>\<close>; | 
| 35741 | 105 | |
| 106 | fun mk_typedef newT oldT RepC AbsC A = | |
| 80631 | 107 | let val type_definition = \<^Const>\<open>type_definition newT oldT for RepC AbsC A\<close> | 
| 108 | in Logic.mk_implies (mk_inhabited oldT A, HOLogic.mk_Trueprop type_definition) end; | |
| 35134 | 109 | |
| 61260 | 110 | fun primitive_typedef {overloaded} type_definition_name newT oldT Rep_name Abs_name A lthy =
 | 
| 35741 | 111 | let | 
| 112 | (* errors *) | |
| 113 | ||
| 114 | fun show_names pairs = commas_quote (map fst pairs); | |
| 115 | ||
| 116 | val lhs_tfrees = Term.add_tfreesT newT []; | |
| 117 | val rhs_tfrees = Term.add_tfreesT oldT []; | |
| 118 | val _ = | |
| 61260 | 119 | (case fold (remove (op =)) lhs_tfrees rhs_tfrees of | 
| 120 | [] => () | |
| 35741 | 121 |       | extras => error ("Extra type variables in representing set: " ^ show_names extras));
 | 
| 122 | ||
| 123 | val _ = | |
| 61260 | 124 | (case Term.add_frees A [] of [] => | 
| 125 | [] | |
| 35741 | 126 |       | xs => error ("Illegal variables in representing set: " ^ show_names xs));
 | 
| 35134 | 127 | |
| 35741 | 128 | |
| 129 | (* axiomatization *) | |
| 130 | ||
| 42375 
774df7c59508
report Name_Space.declare/define, relatively to context;
 wenzelm parents: 
42361diff
changeset | 131 | val ((RepC, AbsC), consts_lthy) = lthy | 
| 
774df7c59508
report Name_Space.declare/define, relatively to context;
 wenzelm parents: 
42361diff
changeset | 132 | |> Local_Theory.background_theory_result | 
| 
774df7c59508
report Name_Space.declare/define, relatively to context;
 wenzelm parents: 
42361diff
changeset | 133 | (Sign.declare_const lthy ((Rep_name, newT --> oldT), NoSyn) ##>> | 
| 
774df7c59508
report Name_Space.declare/define, relatively to context;
 wenzelm parents: 
42361diff
changeset | 134 | Sign.declare_const lthy ((Abs_name, oldT --> newT), NoSyn)); | 
| 61255 
15865e0c5598
eliminated separate type Theory.dep: use typeargs uniformly for consts/types;
 wenzelm parents: 
61248diff
changeset | 135 | val const_dep = Theory.const_dep (Proof_Context.theory_of consts_lthy); | 
| 61261 
ddb2da7cb2e4
more explicit Defs.context: use proper name spaces as far as possible;
 wenzelm parents: 
61260diff
changeset | 136 | val defs_context = Proof_Context.defs_context consts_lthy; | 
| 61247 | 137 | |
| 61255 
15865e0c5598
eliminated separate type Theory.dep: use typeargs uniformly for consts/types;
 wenzelm parents: 
61248diff
changeset | 138 | val A_consts = fold_aterms (fn Const c => insert (op =) (const_dep c) | _ => I) A []; | 
| 61248 | 139 | val A_types = | 
| 61255 
15865e0c5598
eliminated separate type Theory.dep: use typeargs uniformly for consts/types;
 wenzelm parents: 
61248diff
changeset | 140 | (fold_types o fold_subtypes) (fn Type t => insert (op =) (Theory.type_dep t) | _ => I) A []; | 
| 61246 
077b88f9ec16
HOL typedef with explicit dependency checks according to Ondrey Kuncar, 07-Jul-2015, 16-Jul-2015, 30-Jul-2015;
 wenzelm parents: 
61110diff
changeset | 141 | val typedef_deps = A_consts @ A_types; | 
| 61247 | 142 | |
| 61255 
15865e0c5598
eliminated separate type Theory.dep: use typeargs uniformly for consts/types;
 wenzelm parents: 
61248diff
changeset | 143 | val newT_dep = Theory.type_dep (dest_Type newT); | 
| 35741 | 144 | |
| 42375 
774df7c59508
report Name_Space.declare/define, relatively to context;
 wenzelm parents: 
42361diff
changeset | 145 | val ((axiom_name, axiom), axiom_lthy) = consts_lthy | 
| 
774df7c59508
report Name_Space.declare/define, relatively to context;
 wenzelm parents: 
42361diff
changeset | 146 | |> Local_Theory.background_theory_result | 
| 61110 | 147 | (Thm.add_axiom consts_lthy (type_definition_name, mk_typedef newT oldT RepC AbsC A) ##> | 
| 61261 
ddb2da7cb2e4
more explicit Defs.context: use proper name spaces as far as possible;
 wenzelm parents: 
61260diff
changeset | 148 | Theory.add_deps defs_context "" newT_dep typedef_deps ##> | 
| 
ddb2da7cb2e4
more explicit Defs.context: use proper name spaces as far as possible;
 wenzelm parents: 
61260diff
changeset | 149 | Theory.add_deps defs_context "" (const_dep (dest_Const RepC)) [newT_dep] ##> | 
| 
ddb2da7cb2e4
more explicit Defs.context: use proper name spaces as far as possible;
 wenzelm parents: 
61260diff
changeset | 150 | Theory.add_deps defs_context "" (const_dep (dest_Const AbsC)) [newT_dep]); | 
| 35741 | 151 | |
| 61260 | 152 | val axiom_defs = Theory.defs_of (Proof_Context.theory_of axiom_lthy); | 
| 153 | val newT_deps = maps #2 (Defs.get_deps axiom_defs (#1 newT_dep)); | |
| 154 | val _ = | |
| 155 | if null newT_deps orelse overloaded orelse Config.get lthy typedef_overloaded then () | |
| 156 | else | |
| 157 | error (Pretty.string_of (Pretty.chunks | |
| 158 | [Pretty.paragraph | |
| 159 | (Pretty.text "Type definition with open dependencies, use" @ | |
| 160 | [Pretty.brk 1, Pretty.str "\"", Pretty.keyword1 "typedef", Pretty.brk 1, | |
| 161 |               Pretty.str "(", Pretty.keyword2 "overloaded", Pretty.str ")\"", Pretty.brk 1] @
 | |
| 162 | Pretty.text "or enable configuration option \"typedef_overloaded\" in the context."), | |
| 163 | Pretty.block [Pretty.str " Type:", Pretty.brk 2, Syntax.pretty_typ axiom_lthy newT], | |
| 164 | Pretty.block (Pretty.str " Deps:" :: Pretty.brk 2 :: | |
| 61261 
ddb2da7cb2e4
more explicit Defs.context: use proper name spaces as far as possible;
 wenzelm parents: 
61260diff
changeset | 165 | Pretty.commas | 
| 
ddb2da7cb2e4
more explicit Defs.context: use proper name spaces as far as possible;
 wenzelm parents: 
61260diff
changeset | 166 | (map (Defs.pretty_entry (Proof_Context.defs_context axiom_lthy)) newT_deps))])) | 
| 42375 
774df7c59508
report Name_Space.declare/define, relatively to context;
 wenzelm parents: 
42361diff
changeset | 167 | in ((RepC, AbsC, axiom_name, axiom), axiom_lthy) end; | 
| 15259 | 168 | |
| 169 | ||
| 61110 | 170 | (* derived bindings *) | 
| 171 | ||
| 172 | type bindings = {Rep_name: binding, Abs_name: binding, type_definition_name: binding};
 | |
| 173 | ||
| 62513 
702085ca8564
take qualification of type name more seriously: derived consts and facts are qualified uniformly;
 wenzelm parents: 
61261diff
changeset | 174 | fun prefix_binding prfx name = | 
| 63003 | 175 | Binding.reset_pos (Binding.qualify_name false name (prfx ^ Binding.name_of name)); | 
| 62513 
702085ca8564
take qualification of type name more seriously: derived consts and facts are qualified uniformly;
 wenzelm parents: 
61261diff
changeset | 176 | |
| 
702085ca8564
take qualification of type name more seriously: derived consts and facts are qualified uniformly;
 wenzelm parents: 
61261diff
changeset | 177 | fun qualify_binding name = Binding.qualify false (Binding.name_of name); | 
| 
702085ca8564
take qualification of type name more seriously: derived consts and facts are qualified uniformly;
 wenzelm parents: 
61261diff
changeset | 178 | |
| 61110 | 179 | fun default_bindings name = | 
| 62513 
702085ca8564
take qualification of type name more seriously: derived consts and facts are qualified uniformly;
 wenzelm parents: 
61261diff
changeset | 180 |  {Rep_name = prefix_binding "Rep_" name,
 | 
| 
702085ca8564
take qualification of type name more seriously: derived consts and facts are qualified uniformly;
 wenzelm parents: 
61261diff
changeset | 181 | Abs_name = prefix_binding "Abs_" name, | 
| 
702085ca8564
take qualification of type name more seriously: derived consts and facts are qualified uniformly;
 wenzelm parents: 
61261diff
changeset | 182 | type_definition_name = prefix_binding "type_definition_" name}; | 
| 61110 | 183 | |
| 184 | fun make_bindings name NONE = default_bindings name | |
| 185 | | make_bindings _ (SOME bindings) = bindings; | |
| 186 | ||
| 187 | fun make_morphisms name NONE = default_bindings name | |
| 188 | | make_morphisms name (SOME (Rep_name, Abs_name)) = | |
| 62513 
702085ca8564
take qualification of type name more seriously: derived consts and facts are qualified uniformly;
 wenzelm parents: 
61261diff
changeset | 189 |      {Rep_name = qualify_binding name Rep_name,
 | 
| 
702085ca8564
take qualification of type name more seriously: derived consts and facts are qualified uniformly;
 wenzelm parents: 
61261diff
changeset | 190 | Abs_name = qualify_binding name Abs_name, | 
| 
702085ca8564
take qualification of type name more seriously: derived consts and facts are qualified uniformly;
 wenzelm parents: 
61261diff
changeset | 191 | type_definition_name = #type_definition_name (default_bindings name)}; | 
| 61110 | 192 | |
| 193 | ||
| 6383 | 194 | (* prepare_typedef *) | 
| 195 | ||
| 61260 | 196 | fun prepare_typedef prep_term overloaded (name, raw_args, mx) raw_set opt_bindings lthy = | 
| 4866 | 197 | let | 
| 35741 | 198 | (* rhs *) | 
| 199 | ||
| 36153 
1ac501e16a6a
replaced slightly odd Typedecl.predeclare_constraints by plain declaration of type arguments -- also avoid "recursive" declaration of type constructor, which can cause problems with sequential definitions B.foo = A.foo;
 wenzelm parents: 
36107diff
changeset | 200 | val tmp_ctxt = lthy |> fold (Variable.declare_typ o TFree) raw_args; | 
| 
1ac501e16a6a
replaced slightly odd Typedecl.predeclare_constraints by plain declaration of type arguments -- also avoid "recursive" declaration of type constructor, which can cause problems with sequential definitions B.foo = A.foo;
 wenzelm parents: 
36107diff
changeset | 201 | val set = prep_term tmp_ctxt raw_set; | 
| 
1ac501e16a6a
replaced slightly odd Typedecl.predeclare_constraints by plain declaration of type arguments -- also avoid "recursive" declaration of type constructor, which can cause problems with sequential definitions B.foo = A.foo;
 wenzelm parents: 
36107diff
changeset | 202 | val tmp_ctxt' = tmp_ctxt |> Variable.declare_term set; | 
| 35836 
9380fab5f4f7
eliminated slightly odd typedecl_wrt in favour of explicit predeclare_constraints;
 wenzelm parents: 
35766diff
changeset | 203 | |
| 21352 | 204 | val setT = Term.fastype_of set; | 
| 35741 | 205 | val oldT = HOLogic.dest_setT setT handle TYPE _ => | 
| 206 |       error ("Not a set type: " ^ quote (Syntax.string_of_typ lthy setT));
 | |
| 207 | ||
| 62513 
702085ca8564
take qualification of type name more seriously: derived consts and facts are qualified uniformly;
 wenzelm parents: 
61261diff
changeset | 208 | val bname = Binding.name_of name; | 
| 80631 | 209 | val goal = mk_inhabited oldT set; | 
| 210 | val goal_pat = | |
| 211 | mk_inhabited oldT (Var (the_default (bname, 0) (Lexicon.read_variable bname), setT)); | |
| 35741 | 212 | |
| 213 | ||
| 214 | (* lhs *) | |
| 215 | ||
| 42361 | 216 | val args = map (Proof_Context.check_tfree tmp_ctxt') raw_args; | 
| 35741 | 217 | val (newT, typedecl_lthy) = lthy | 
| 61259 | 218 |       |> Typedecl.typedecl {final = false} (name, args, mx)
 | 
| 35741 | 219 | ||> Variable.declare_term set; | 
| 220 | ||
| 80633 | 221 | val full_name = dest_Type_name newT; | 
| 35741 | 222 | |
| 223 | ||
| 224 | (* axiomatization *) | |
| 4866 | 225 | |
| 61110 | 226 |     val {Rep_name, Abs_name, type_definition_name} = make_bindings name opt_bindings;
 | 
| 4866 | 227 | |
| 49833 | 228 | val ((RepC, AbsC, axiom_name, typedef), typedef_lthy) = typedecl_lthy | 
| 61260 | 229 | |> primitive_typedef overloaded type_definition_name newT oldT Rep_name Abs_name set; | 
| 35741 | 230 | |
| 231 | val alias_lthy = typedef_lthy | |
| 80636 
4041e7c8059d
tuned: more explicit dest_Const_name and dest_Const_type;
 wenzelm parents: 
80633diff
changeset | 232 | |> Local_Theory.const_alias Rep_name (Term.dest_Const_name RepC) | 
| 
4041e7c8059d
tuned: more explicit dest_Const_name and dest_Const_type;
 wenzelm parents: 
80633diff
changeset | 233 | |> Local_Theory.const_alias Abs_name (Term.dest_Const_name AbsC); | 
| 6383 | 234 | |
| 29056 | 235 | |
| 35741 | 236 | (* result *) | 
| 4866 | 237 | |
| 62513 
702085ca8564
take qualification of type name more seriously: derived consts and facts are qualified uniformly;
 wenzelm parents: 
61261diff
changeset | 238 | fun note ((b, atts), th) = | 
| 63019 | 239 | Local_Theory.note ((b, atts), [th]) #>> (fn (_, [th']) => th'); | 
| 4866 | 240 | |
| 35741 | 241 | fun typedef_result inhabited lthy1 = | 
| 242 | let | |
| 70483 | 243 | val ((_, [type_definition]), lthy2) = lthy1 | 
| 244 | |> Local_Theory.note ((type_definition_name, []), [inhabited RS typedef]); | |
| 245 | fun make th = Goal.norm_result lthy2 (type_definition RS th); | |
| 246 | val (((((((((Rep, Rep_inverse), Abs_inverse), Rep_inject), Abs_inject), Rep_cases), | |
| 247 | Abs_cases), Rep_induct), Abs_induct), lthy3) = lthy2 | |
| 248 |           |> note ((Rep_name, []), make @{thm type_definition.Rep})
 | |
| 62513 
702085ca8564
take qualification of type name more seriously: derived consts and facts are qualified uniformly;
 wenzelm parents: 
61261diff
changeset | 249 | ||>> note ((Binding.suffix_name "_inverse" Rep_name, []), | 
| 35741 | 250 |               make @{thm type_definition.Rep_inverse})
 | 
| 62513 
702085ca8564
take qualification of type name more seriously: derived consts and facts are qualified uniformly;
 wenzelm parents: 
61261diff
changeset | 251 | ||>> note ((Binding.suffix_name "_inverse" Abs_name, []), | 
| 35741 | 252 |               make @{thm type_definition.Abs_inverse})
 | 
| 62513 
702085ca8564
take qualification of type name more seriously: derived consts and facts are qualified uniformly;
 wenzelm parents: 
61261diff
changeset | 253 | ||>> note ((Binding.suffix_name "_inject" Rep_name, []), | 
| 35741 | 254 |               make @{thm type_definition.Rep_inject})
 | 
| 62513 
702085ca8564
take qualification of type name more seriously: derived consts and facts are qualified uniformly;
 wenzelm parents: 
61261diff
changeset | 255 | ||>> note ((Binding.suffix_name "_inject" Abs_name, []), | 
| 35741 | 256 |               make @{thm type_definition.Abs_inject})
 | 
| 62513 
702085ca8564
take qualification of type name more seriously: derived consts and facts are qualified uniformly;
 wenzelm parents: 
61261diff
changeset | 257 | ||>> note ((Binding.suffix_name "_cases" Rep_name, | 
| 63019 | 258 | [Attrib.case_names [Binding.name_of Rep_name], | 
| 78095 | 259 | Attrib.internal \<^here> (K (Induct.cases_pred full_name))]), | 
| 35741 | 260 |               make @{thm type_definition.Rep_cases})
 | 
| 62513 
702085ca8564
take qualification of type name more seriously: derived consts and facts are qualified uniformly;
 wenzelm parents: 
61261diff
changeset | 261 | ||>> note ((Binding.suffix_name "_cases" Abs_name, | 
| 63019 | 262 | [Attrib.case_names [Binding.name_of Abs_name], | 
| 78095 | 263 | Attrib.internal \<^here> (K (Induct.cases_type full_name))]), | 
| 35741 | 264 |               make @{thm type_definition.Abs_cases})
 | 
| 62513 
702085ca8564
take qualification of type name more seriously: derived consts and facts are qualified uniformly;
 wenzelm parents: 
61261diff
changeset | 265 | ||>> note ((Binding.suffix_name "_induct" Rep_name, | 
| 63019 | 266 | [Attrib.case_names [Binding.name_of Rep_name], | 
| 78095 | 267 | Attrib.internal \<^here> (K (Induct.induct_pred full_name))]), | 
| 35741 | 268 |               make @{thm type_definition.Rep_induct})
 | 
| 62513 
702085ca8564
take qualification of type name more seriously: derived consts and facts are qualified uniformly;
 wenzelm parents: 
61261diff
changeset | 269 | ||>> note ((Binding.suffix_name "_induct" Abs_name, | 
| 63019 | 270 | [Attrib.case_names [Binding.name_of Abs_name], | 
| 78095 | 271 | Attrib.internal \<^here> (K (Induct.induct_type full_name))]), | 
| 35741 | 272 |               make @{thm type_definition.Abs_induct});
 | 
| 4866 | 273 | |
| 35994 
9cc3df9a606e
Typedef.info: separate global and local part, only the latter is transformed by morphisms;
 wenzelm parents: 
35840diff
changeset | 274 | val info = | 
| 80636 
4041e7c8059d
tuned: more explicit dest_Const_name and dest_Const_type;
 wenzelm parents: 
80633diff
changeset | 275 |           ({rep_type = oldT, abs_type = newT, Rep_name = dest_Const_name RepC,
 | 
| 
4041e7c8059d
tuned: more explicit dest_Const_name and dest_Const_type;
 wenzelm parents: 
80633diff
changeset | 276 | Abs_name = dest_Const_name AbsC, axiom_name = axiom_name}, | 
| 49833 | 277 |            {inhabited = inhabited, type_definition = type_definition,
 | 
| 35741 | 278 | Rep = Rep, Rep_inverse = Rep_inverse, Abs_inverse = Abs_inverse, | 
| 279 | Rep_inject = Rep_inject, Abs_inject = Abs_inject, Rep_cases = Rep_cases, | |
| 35994 
9cc3df9a606e
Typedef.info: separate global and local part, only the latter is transformed by morphisms;
 wenzelm parents: 
35840diff
changeset | 280 | Abs_cases = Abs_cases, Rep_induct = Rep_induct, Abs_induct = Abs_induct}); | 
| 35741 | 281 | in | 
| 70483 | 282 | lthy3 | 
| 78095 | 283 |         |> Local_Theory.declaration {syntax = false, pervasive = true, pos = \<^here>}
 | 
| 49835 | 284 | (fn phi => put_info full_name (transform_info phi info)) | 
| 58662 | 285 | |> Typedef_Plugin.data Plugin_Name.default_filter full_name | 
| 49835 | 286 | |> pair (full_name, info) | 
| 35741 | 287 | end; | 
| 11426 
f280d4b29a2c
abtract non-emptiness statements (no longer use Eps);
 wenzelm parents: 
10697diff
changeset | 288 | |
| 35741 | 289 | in ((goal, goal_pat, typedef_result), alias_lthy) end | 
| 30345 | 290 | handle ERROR msg => | 
| 42381 
309ec68442c6
added Binding.print convenience, which includes quote already;
 wenzelm parents: 
42375diff
changeset | 291 |     cat_error msg ("The error(s) above occurred in typedef " ^ Binding.print name);
 | 
| 4866 | 292 | |
| 293 | ||
| 29056 | 294 | (* add_typedef: tactic interface *) | 
| 4866 | 295 | |
| 61260 | 296 | fun add_typedef overloaded typ set opt_bindings tac lthy = | 
| 6383 | 297 | let | 
| 35741 | 298 | val ((goal, _, typedef_result), lthy') = | 
| 61260 | 299 | prepare_typedef Syntax.check_term overloaded typ set opt_bindings lthy; | 
| 70483 | 300 | val inhabited = Goal.prove lthy' [] [] goal (tac o #context) |> Goal.norm_result lthy'; | 
| 35741 | 301 | in typedef_result inhabited lthy' end; | 
| 302 | ||
| 61260 | 303 | fun add_typedef_global overloaded typ set opt_bindings tac = | 
| 69829 | 304 | Named_Target.theory_map_result (apsnd o transform_info) | 
| 305 | (add_typedef overloaded typ set opt_bindings tac) | |
| 4866 | 306 | |
| 17339 | 307 | |
| 29056 | 308 | (* typedef: proof interface *) | 
| 6383 | 309 | |
| 17339 | 310 | local | 
| 311 | ||
| 61260 | 312 | fun gen_typedef prep_term prep_constraint overloaded (b, raw_args, mx) set opt_bindings lthy = | 
| 11822 | 313 | let | 
| 35840 
01d7c4ba9050
allow sort constraints in HOL/typedef and related HOLCF variants;
 wenzelm parents: 
35836diff
changeset | 314 | val args = map (apsnd (prep_constraint lthy)) raw_args; | 
| 35741 | 315 | val ((goal, goal_pat, typedef_result), lthy') = | 
| 61260 | 316 | prepare_typedef prep_term overloaded (b, args, mx) set opt_bindings lthy; | 
| 35741 | 317 | fun after_qed [[th]] = snd o typedef_result th; | 
| 36323 
655e2d74de3a
modernized naming conventions of main Isar proof elements;
 wenzelm parents: 
36153diff
changeset | 318 | in Proof.theorem NONE after_qed [[(goal, [goal_pat])]] lthy' end; | 
| 17339 | 319 | |
| 320 | in | |
| 6383 | 321 | |
| 35840 
01d7c4ba9050
allow sort constraints in HOL/typedef and related HOLCF variants;
 wenzelm parents: 
35836diff
changeset | 322 | val typedef = gen_typedef Syntax.check_term (K I); | 
| 
01d7c4ba9050
allow sort constraints in HOL/typedef and related HOLCF variants;
 wenzelm parents: 
35836diff
changeset | 323 | val typedef_cmd = gen_typedef Syntax.read_term Typedecl.read_constraint; | 
| 17339 | 324 | |
| 19705 | 325 | end; | 
| 15259 | 326 | |
| 327 | ||
| 328 | ||
| 6383 | 329 | (** outer syntax **) | 
| 330 | ||
| 24867 | 331 | val _ = | 
| 67149 | 332 | Outer_Syntax.local_theory_to_proof \<^command_keyword>\<open>typedef\<close> | 
| 46961 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 wenzelm parents: 
46949diff
changeset | 333 | "HOL type definition (requires non-emptiness proof)" | 
| 61260 | 334 | (Parse_Spec.overloaded -- Parse.type_args_constrained -- Parse.binding -- Parse.opt_mixfix -- | 
| 67149 | 335 | (\<^keyword>\<open>=\<close> |-- Parse.term) -- | 
| 336 | Scan.option (\<^keyword>\<open>morphisms\<close> |-- Parse.!!! (Parse.binding -- Parse.binding)) | |
| 61260 | 337 | >> (fn (((((overloaded, vs), t), mx), A), opt_morphs) => fn lthy => | 
| 338 |         typedef_cmd {overloaded = overloaded} (t, vs, mx) A
 | |
| 339 | (SOME (make_morphisms t opt_morphs)) lthy)); | |
| 340 | ||
| 341 | ||
| 342 | val overloaded = typedef_overloaded; | |
| 6357 | 343 | |
| 68264 | 344 | |
| 345 | ||
| 346 | (** theory export **) | |
| 347 | ||
| 74114 
700e5bd59c7d
clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
 wenzelm parents: 
70483diff
changeset | 348 | val _ = | 
| 
700e5bd59c7d
clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
 wenzelm parents: 
70483diff
changeset | 349 | (Theory.setup o Thy_Info.add_presentation) (fn context => fn thy => | 
| 
700e5bd59c7d
clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
 wenzelm parents: 
70483diff
changeset | 350 | if Export_Theory.export_enabled context then | 
| 
700e5bd59c7d
clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
 wenzelm parents: 
70483diff
changeset | 351 | let | 
| 
700e5bd59c7d
clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
 wenzelm parents: 
70483diff
changeset | 352 | val parent_spaces = map Sign.type_space (Theory.parents_of thy); | 
| 
700e5bd59c7d
clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
 wenzelm parents: 
70483diff
changeset | 353 | val typedefs = | 
| 
700e5bd59c7d
clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
 wenzelm parents: 
70483diff
changeset | 354 | Name_Space.dest_table (#types (Type.rep_tsig (Sign.tsig_of thy))) | 
| 
700e5bd59c7d
clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
 wenzelm parents: 
70483diff
changeset | 355 | |> maps (fn (name, _) => | 
| 
700e5bd59c7d
clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
 wenzelm parents: 
70483diff
changeset | 356 | if exists (fn space => Name_Space.declared space name) parent_spaces then [] | 
| 
700e5bd59c7d
clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
 wenzelm parents: 
70483diff
changeset | 357 | else | 
| 
700e5bd59c7d
clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
 wenzelm parents: 
70483diff
changeset | 358 | get_info_global thy name | 
| 
700e5bd59c7d
clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
 wenzelm parents: 
70483diff
changeset | 359 |                 |> map (fn ({rep_type, abs_type, Rep_name, Abs_name, axiom_name}, _) =>
 | 
| 
700e5bd59c7d
clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
 wenzelm parents: 
70483diff
changeset | 360 | (name, (rep_type, (abs_type, (Rep_name, (Abs_name, axiom_name))))))); | 
| 
700e5bd59c7d
clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
 wenzelm parents: 
70483diff
changeset | 361 | val encode = | 
| 
700e5bd59c7d
clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
 wenzelm parents: 
70483diff
changeset | 362 | let open XML.Encode Term_XML.Encode | 
| 
700e5bd59c7d
clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
 wenzelm parents: 
70483diff
changeset | 363 | in list (pair string (pair typ (pair typ (pair string (pair string string))))) end; | 
| 
700e5bd59c7d
clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
 wenzelm parents: 
70483diff
changeset | 364 | in | 
| 
700e5bd59c7d
clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
 wenzelm parents: 
70483diff
changeset | 365 | if null typedefs then () | 
| 
700e5bd59c7d
clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
 wenzelm parents: 
70483diff
changeset | 366 | else Export_Theory.export_body thy "typedefs" (encode typedefs) | 
| 
700e5bd59c7d
clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
 wenzelm parents: 
70483diff
changeset | 367 | end | 
| 
700e5bd59c7d
clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
 wenzelm parents: 
70483diff
changeset | 368 | else ()); | 
| 68264 | 369 | |
| 29056 | 370 | end; |