author | wenzelm |
Sun, 29 Nov 2020 17:54:50 +0100 | |
changeset 72781 | 15a8de807f21 |
parent 70483 | 59250a328113 |
child 74114 | 700e5bd59c7d |
permissions | -rw-r--r-- |
31723
f5cafe803b55
discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents:
30345
diff
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:
30345
diff
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:
29059
diff
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:
30345
diff
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:
35840
diff
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:
35840
diff
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:
35840
diff
changeset
|
57 |
(global_info, |
9cc3df9a606e
Typedef.info: separate global and local part, only the latter is transformed by morphisms;
wenzelm
parents:
35840
diff
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; |
16458 | 69 |
val extend = I; |
35741 | 70 |
fun merge data = Symtab.merge_list (K true) data; |
22846 | 71 |
); |
15259 | 72 |
|
61103 | 73 |
fun get_info_generic context = |
74 |
Symtab.lookup_list (Data.get context) #> |
|
67664 | 75 |
map (transform_info (Morphism.transfer_morphism'' context)); |
35741 | 76 |
|
61103 | 77 |
val get_info = get_info_generic o Context.Proof; |
78 |
val get_info_global = get_info_generic o Context.Theory; |
|
79 |
||
80 |
fun put_info name info = |
|
81 |
Data.map (Symtab.cons_list (name, transform_info Morphism.trim_context_morphism info)); |
|
35741 | 82 |
|
83 |
||
84 |
(* global interpretation *) |
|
85 |
||
58662 | 86 |
structure Typedef_Plugin = Plugin(type T = string); |
87 |
||
67149 | 88 |
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:
54883
diff
changeset
|
89 |
|
58662 | 90 |
fun interpretation f = |
91 |
Typedef_Plugin.interpretation typedef_plugin |
|
92 |
(fn name => fn lthy => |
|
93 |
lthy |
|
59880
30687c3f2b10
clarified role of naming for background theory: transform_binding (e.g. for "concealed" flag) uses naming of hypothetical context;
wenzelm
parents:
59859
diff
changeset
|
94 |
|> Local_Theory.map_background_naming |
58662 | 95 |
(Name_Space.root_path #> Name_Space.add_path (Long_Name.qualifier name)) |
96 |
|> 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:
59859
diff
changeset
|
97 |
|> Local_Theory.restore_background_naming lthy); |
35741 | 98 |
|
99 |
||
100 |
(* primitive typedef axiomatization -- for fresh typedecl *) |
|
101 |
||
67149 | 102 |
val typedef_overloaded = Attrib.setup_config_bool \<^binding>\<open>typedef_overloaded\<close> (K false); |
61260 | 103 |
|
35741 | 104 |
fun mk_inhabited A = |
105 |
let val T = HOLogic.dest_setT (Term.fastype_of A) |
|
106 |
in HOLogic.mk_Trueprop (HOLogic.exists_const T $ Abs ("x", T, HOLogic.mk_mem (Bound 0, A))) end; |
|
107 |
||
108 |
fun mk_typedef newT oldT RepC AbsC A = |
|
109 |
let |
|
110 |
val typedefC = |
|
67149 | 111 |
Const (\<^const_name>\<open>type_definition\<close>, |
35741 | 112 |
(newT --> oldT) --> (oldT --> newT) --> HOLogic.mk_setT oldT --> HOLogic.boolT); |
113 |
in Logic.mk_implies (mk_inhabited A, HOLogic.mk_Trueprop (typedefC $ RepC $ AbsC $ A)) end; |
|
35134 | 114 |
|
61260 | 115 |
fun primitive_typedef {overloaded} type_definition_name newT oldT Rep_name Abs_name A lthy = |
35741 | 116 |
let |
117 |
(* errors *) |
|
118 |
||
119 |
fun show_names pairs = commas_quote (map fst pairs); |
|
120 |
||
121 |
val lhs_tfrees = Term.add_tfreesT newT []; |
|
122 |
val rhs_tfrees = Term.add_tfreesT oldT []; |
|
123 |
val _ = |
|
61260 | 124 |
(case fold (remove (op =)) lhs_tfrees rhs_tfrees of |
125 |
[] => () |
|
35741 | 126 |
| extras => error ("Extra type variables in representing set: " ^ show_names extras)); |
127 |
||
128 |
val _ = |
|
61260 | 129 |
(case Term.add_frees A [] of [] => |
130 |
[] |
|
35741 | 131 |
| xs => error ("Illegal variables in representing set: " ^ show_names xs)); |
35134 | 132 |
|
35741 | 133 |
|
134 |
(* axiomatization *) |
|
135 |
||
42375
774df7c59508
report Name_Space.declare/define, relatively to context;
wenzelm
parents:
42361
diff
changeset
|
136 |
val ((RepC, AbsC), consts_lthy) = lthy |
774df7c59508
report Name_Space.declare/define, relatively to context;
wenzelm
parents:
42361
diff
changeset
|
137 |
|> Local_Theory.background_theory_result |
774df7c59508
report Name_Space.declare/define, relatively to context;
wenzelm
parents:
42361
diff
changeset
|
138 |
(Sign.declare_const lthy ((Rep_name, newT --> oldT), NoSyn) ##>> |
774df7c59508
report Name_Space.declare/define, relatively to context;
wenzelm
parents:
42361
diff
changeset
|
139 |
Sign.declare_const lthy ((Abs_name, oldT --> newT), NoSyn)); |
61255
15865e0c5598
eliminated separate type Theory.dep: use typeargs uniformly for consts/types;
wenzelm
parents:
61248
diff
changeset
|
140 |
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:
61260
diff
changeset
|
141 |
val defs_context = Proof_Context.defs_context consts_lthy; |
61247 | 142 |
|
61255
15865e0c5598
eliminated separate type Theory.dep: use typeargs uniformly for consts/types;
wenzelm
parents:
61248
diff
changeset
|
143 |
val A_consts = fold_aterms (fn Const c => insert (op =) (const_dep c) | _ => I) A []; |
61248 | 144 |
val A_types = |
61255
15865e0c5598
eliminated separate type Theory.dep: use typeargs uniformly for consts/types;
wenzelm
parents:
61248
diff
changeset
|
145 |
(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:
61110
diff
changeset
|
146 |
val typedef_deps = A_consts @ A_types; |
61247 | 147 |
|
61255
15865e0c5598
eliminated separate type Theory.dep: use typeargs uniformly for consts/types;
wenzelm
parents:
61248
diff
changeset
|
148 |
val newT_dep = Theory.type_dep (dest_Type newT); |
35741 | 149 |
|
42375
774df7c59508
report Name_Space.declare/define, relatively to context;
wenzelm
parents:
42361
diff
changeset
|
150 |
val ((axiom_name, axiom), axiom_lthy) = consts_lthy |
774df7c59508
report Name_Space.declare/define, relatively to context;
wenzelm
parents:
42361
diff
changeset
|
151 |
|> Local_Theory.background_theory_result |
61110 | 152 |
(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:
61260
diff
changeset
|
153 |
Theory.add_deps defs_context "" newT_dep typedef_deps ##> |
ddb2da7cb2e4
more explicit Defs.context: use proper name spaces as far as possible;
wenzelm
parents:
61260
diff
changeset
|
154 |
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:
61260
diff
changeset
|
155 |
Theory.add_deps defs_context "" (const_dep (dest_Const AbsC)) [newT_dep]); |
35741 | 156 |
|
61260 | 157 |
val axiom_defs = Theory.defs_of (Proof_Context.theory_of axiom_lthy); |
158 |
val newT_deps = maps #2 (Defs.get_deps axiom_defs (#1 newT_dep)); |
|
159 |
val _ = |
|
160 |
if null newT_deps orelse overloaded orelse Config.get lthy typedef_overloaded then () |
|
161 |
else |
|
162 |
error (Pretty.string_of (Pretty.chunks |
|
163 |
[Pretty.paragraph |
|
164 |
(Pretty.text "Type definition with open dependencies, use" @ |
|
165 |
[Pretty.brk 1, Pretty.str "\"", Pretty.keyword1 "typedef", Pretty.brk 1, |
|
166 |
Pretty.str "(", Pretty.keyword2 "overloaded", Pretty.str ")\"", Pretty.brk 1] @ |
|
167 |
Pretty.text "or enable configuration option \"typedef_overloaded\" in the context."), |
|
168 |
Pretty.block [Pretty.str " Type:", Pretty.brk 2, Syntax.pretty_typ axiom_lthy newT], |
|
169 |
Pretty.block (Pretty.str " Deps:" :: Pretty.brk 2 :: |
|
61261
ddb2da7cb2e4
more explicit Defs.context: use proper name spaces as far as possible;
wenzelm
parents:
61260
diff
changeset
|
170 |
Pretty.commas |
ddb2da7cb2e4
more explicit Defs.context: use proper name spaces as far as possible;
wenzelm
parents:
61260
diff
changeset
|
171 |
(map (Defs.pretty_entry (Proof_Context.defs_context axiom_lthy)) newT_deps))])) |
42375
774df7c59508
report Name_Space.declare/define, relatively to context;
wenzelm
parents:
42361
diff
changeset
|
172 |
in ((RepC, AbsC, axiom_name, axiom), axiom_lthy) end; |
15259 | 173 |
|
174 |
||
61110 | 175 |
(* derived bindings *) |
176 |
||
177 |
type bindings = {Rep_name: binding, Abs_name: binding, type_definition_name: binding}; |
|
178 |
||
62513
702085ca8564
take qualification of type name more seriously: derived consts and facts are qualified uniformly;
wenzelm
parents:
61261
diff
changeset
|
179 |
fun prefix_binding prfx name = |
63003 | 180 |
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:
61261
diff
changeset
|
181 |
|
702085ca8564
take qualification of type name more seriously: derived consts and facts are qualified uniformly;
wenzelm
parents:
61261
diff
changeset
|
182 |
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:
61261
diff
changeset
|
183 |
|
61110 | 184 |
fun default_bindings name = |
62513
702085ca8564
take qualification of type name more seriously: derived consts and facts are qualified uniformly;
wenzelm
parents:
61261
diff
changeset
|
185 |
{Rep_name = prefix_binding "Rep_" name, |
702085ca8564
take qualification of type name more seriously: derived consts and facts are qualified uniformly;
wenzelm
parents:
61261
diff
changeset
|
186 |
Abs_name = prefix_binding "Abs_" name, |
702085ca8564
take qualification of type name more seriously: derived consts and facts are qualified uniformly;
wenzelm
parents:
61261
diff
changeset
|
187 |
type_definition_name = prefix_binding "type_definition_" name}; |
61110 | 188 |
|
189 |
fun make_bindings name NONE = default_bindings name |
|
190 |
| make_bindings _ (SOME bindings) = bindings; |
|
191 |
||
192 |
fun make_morphisms name NONE = default_bindings name |
|
193 |
| 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:
61261
diff
changeset
|
194 |
{Rep_name = qualify_binding name Rep_name, |
702085ca8564
take qualification of type name more seriously: derived consts and facts are qualified uniformly;
wenzelm
parents:
61261
diff
changeset
|
195 |
Abs_name = qualify_binding name Abs_name, |
702085ca8564
take qualification of type name more seriously: derived consts and facts are qualified uniformly;
wenzelm
parents:
61261
diff
changeset
|
196 |
type_definition_name = #type_definition_name (default_bindings name)}; |
61110 | 197 |
|
198 |
||
6383 | 199 |
(* prepare_typedef *) |
200 |
||
61260 | 201 |
fun prepare_typedef prep_term overloaded (name, raw_args, mx) raw_set opt_bindings lthy = |
4866 | 202 |
let |
35741 | 203 |
(* rhs *) |
204 |
||
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:
36107
diff
changeset
|
205 |
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:
36107
diff
changeset
|
206 |
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:
36107
diff
changeset
|
207 |
val tmp_ctxt' = tmp_ctxt |> Variable.declare_term set; |
35836
9380fab5f4f7
eliminated slightly odd typedecl_wrt in favour of explicit predeclare_constraints;
wenzelm
parents:
35766
diff
changeset
|
208 |
|
21352 | 209 |
val setT = Term.fastype_of set; |
35741 | 210 |
val oldT = HOLogic.dest_setT setT handle TYPE _ => |
211 |
error ("Not a set type: " ^ quote (Syntax.string_of_typ lthy setT)); |
|
212 |
||
62513
702085ca8564
take qualification of type name more seriously: derived consts and facts are qualified uniformly;
wenzelm
parents:
61261
diff
changeset
|
213 |
val bname = Binding.name_of name; |
35741 | 214 |
val goal = mk_inhabited set; |
42290
b1f544c84040
discontinued special treatment of structure Lexicon;
wenzelm
parents:
38757
diff
changeset
|
215 |
val goal_pat = mk_inhabited (Var (the_default (bname, 0) (Lexicon.read_variable bname), setT)); |
35741 | 216 |
|
217 |
||
218 |
(* lhs *) |
|
219 |
||
42361 | 220 |
val args = map (Proof_Context.check_tfree tmp_ctxt') raw_args; |
35741 | 221 |
val (newT, typedecl_lthy) = lthy |
61259 | 222 |
|> Typedecl.typedecl {final = false} (name, args, mx) |
35741 | 223 |
||> Variable.declare_term set; |
224 |
||
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:
54883
diff
changeset
|
225 |
val Type (full_name, _) = newT; |
35741 | 226 |
|
227 |
||
228 |
(* axiomatization *) |
|
4866 | 229 |
|
61110 | 230 |
val {Rep_name, Abs_name, type_definition_name} = make_bindings name opt_bindings; |
4866 | 231 |
|
49833 | 232 |
val ((RepC, AbsC, axiom_name, typedef), typedef_lthy) = typedecl_lthy |
61260 | 233 |
|> primitive_typedef overloaded type_definition_name newT oldT Rep_name Abs_name set; |
35741 | 234 |
|
235 |
val alias_lthy = typedef_lthy |
|
236 |
|> Local_Theory.const_alias Rep_name (#1 (Term.dest_Const RepC)) |
|
237 |
|> Local_Theory.const_alias Abs_name (#1 (Term.dest_Const AbsC)); |
|
6383 | 238 |
|
29056 | 239 |
|
35741 | 240 |
(* result *) |
4866 | 241 |
|
62513
702085ca8564
take qualification of type name more seriously: derived consts and facts are qualified uniformly;
wenzelm
parents:
61261
diff
changeset
|
242 |
fun note ((b, atts), th) = |
63019 | 243 |
Local_Theory.note ((b, atts), [th]) #>> (fn (_, [th']) => th'); |
4866 | 244 |
|
35741 | 245 |
fun typedef_result inhabited lthy1 = |
246 |
let |
|
70483 | 247 |
val ((_, [type_definition]), lthy2) = lthy1 |
248 |
|> Local_Theory.note ((type_definition_name, []), [inhabited RS typedef]); |
|
249 |
fun make th = Goal.norm_result lthy2 (type_definition RS th); |
|
250 |
val (((((((((Rep, Rep_inverse), Abs_inverse), Rep_inject), Abs_inject), Rep_cases), |
|
251 |
Abs_cases), Rep_induct), Abs_induct), lthy3) = lthy2 |
|
252 |
|> 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:
61261
diff
changeset
|
253 |
||>> note ((Binding.suffix_name "_inverse" Rep_name, []), |
35741 | 254 |
make @{thm type_definition.Rep_inverse}) |
62513
702085ca8564
take qualification of type name more seriously: derived consts and facts are qualified uniformly;
wenzelm
parents:
61261
diff
changeset
|
255 |
||>> note ((Binding.suffix_name "_inverse" Abs_name, []), |
35741 | 256 |
make @{thm type_definition.Abs_inverse}) |
62513
702085ca8564
take qualification of type name more seriously: derived consts and facts are qualified uniformly;
wenzelm
parents:
61261
diff
changeset
|
257 |
||>> note ((Binding.suffix_name "_inject" Rep_name, []), |
35741 | 258 |
make @{thm type_definition.Rep_inject}) |
62513
702085ca8564
take qualification of type name more seriously: derived consts and facts are qualified uniformly;
wenzelm
parents:
61261
diff
changeset
|
259 |
||>> note ((Binding.suffix_name "_inject" Abs_name, []), |
35741 | 260 |
make @{thm type_definition.Abs_inject}) |
62513
702085ca8564
take qualification of type name more seriously: derived consts and facts are qualified uniformly;
wenzelm
parents:
61261
diff
changeset
|
261 |
||>> note ((Binding.suffix_name "_cases" Rep_name, |
63019 | 262 |
[Attrib.case_names [Binding.name_of Rep_name], |
263 |
Attrib.internal (K (Induct.cases_pred full_name))]), |
|
35741 | 264 |
make @{thm type_definition.Rep_cases}) |
62513
702085ca8564
take qualification of type name more seriously: derived consts and facts are qualified uniformly;
wenzelm
parents:
61261
diff
changeset
|
265 |
||>> note ((Binding.suffix_name "_cases" Abs_name, |
63019 | 266 |
[Attrib.case_names [Binding.name_of Abs_name], |
267 |
Attrib.internal (K (Induct.cases_type full_name))]), |
|
35741 | 268 |
make @{thm type_definition.Abs_cases}) |
62513
702085ca8564
take qualification of type name more seriously: derived consts and facts are qualified uniformly;
wenzelm
parents:
61261
diff
changeset
|
269 |
||>> note ((Binding.suffix_name "_induct" Rep_name, |
63019 | 270 |
[Attrib.case_names [Binding.name_of Rep_name], |
271 |
Attrib.internal (K (Induct.induct_pred full_name))]), |
|
35741 | 272 |
make @{thm type_definition.Rep_induct}) |
62513
702085ca8564
take qualification of type name more seriously: derived consts and facts are qualified uniformly;
wenzelm
parents:
61261
diff
changeset
|
273 |
||>> note ((Binding.suffix_name "_induct" Abs_name, |
63019 | 274 |
[Attrib.case_names [Binding.name_of Abs_name], |
275 |
Attrib.internal (K (Induct.induct_type full_name))]), |
|
35741 | 276 |
make @{thm type_definition.Abs_induct}); |
4866 | 277 |
|
35994
9cc3df9a606e
Typedef.info: separate global and local part, only the latter is transformed by morphisms;
wenzelm
parents:
35840
diff
changeset
|
278 |
val info = |
36107 | 279 |
({rep_type = oldT, abs_type = newT, Rep_name = #1 (Term.dest_Const RepC), |
280 |
Abs_name = #1 (Term.dest_Const AbsC), axiom_name = axiom_name}, |
|
49833 | 281 |
{inhabited = inhabited, type_definition = type_definition, |
35741 | 282 |
Rep = Rep, Rep_inverse = Rep_inverse, Abs_inverse = Abs_inverse, |
283 |
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:
35840
diff
changeset
|
284 |
Abs_cases = Abs_cases, Rep_induct = Rep_induct, Abs_induct = Abs_induct}); |
35741 | 285 |
in |
70483 | 286 |
lthy3 |
45291
57cd50f98fdc
uniform Local_Theory.declaration with explicit params;
wenzelm
parents:
42381
diff
changeset
|
287 |
|> Local_Theory.declaration {syntax = false, pervasive = true} |
49835 | 288 |
(fn phi => put_info full_name (transform_info phi info)) |
58662 | 289 |
|> Typedef_Plugin.data Plugin_Name.default_filter full_name |
49835 | 290 |
|> pair (full_name, info) |
35741 | 291 |
end; |
11426
f280d4b29a2c
abtract non-emptiness statements (no longer use Eps);
wenzelm
parents:
10697
diff
changeset
|
292 |
|
35741 | 293 |
in ((goal, goal_pat, typedef_result), alias_lthy) end |
30345 | 294 |
handle ERROR msg => |
42381
309ec68442c6
added Binding.print convenience, which includes quote already;
wenzelm
parents:
42375
diff
changeset
|
295 |
cat_error msg ("The error(s) above occurred in typedef " ^ Binding.print name); |
4866 | 296 |
|
297 |
||
29056 | 298 |
(* add_typedef: tactic interface *) |
4866 | 299 |
|
61260 | 300 |
fun add_typedef overloaded typ set opt_bindings tac lthy = |
6383 | 301 |
let |
35741 | 302 |
val ((goal, _, typedef_result), lthy') = |
61260 | 303 |
prepare_typedef Syntax.check_term overloaded typ set opt_bindings lthy; |
70483 | 304 |
val inhabited = Goal.prove lthy' [] [] goal (tac o #context) |> Goal.norm_result lthy'; |
35741 | 305 |
in typedef_result inhabited lthy' end; |
306 |
||
61260 | 307 |
fun add_typedef_global overloaded typ set opt_bindings tac = |
69829 | 308 |
Named_Target.theory_map_result (apsnd o transform_info) |
309 |
(add_typedef overloaded typ set opt_bindings tac) |
|
4866 | 310 |
|
17339 | 311 |
|
29056 | 312 |
(* typedef: proof interface *) |
6383 | 313 |
|
17339 | 314 |
local |
315 |
||
61260 | 316 |
fun gen_typedef prep_term prep_constraint overloaded (b, raw_args, mx) set opt_bindings lthy = |
11822 | 317 |
let |
35840
01d7c4ba9050
allow sort constraints in HOL/typedef and related HOLCF variants;
wenzelm
parents:
35836
diff
changeset
|
318 |
val args = map (apsnd (prep_constraint lthy)) raw_args; |
35741 | 319 |
val ((goal, goal_pat, typedef_result), lthy') = |
61260 | 320 |
prepare_typedef prep_term overloaded (b, args, mx) set opt_bindings lthy; |
35741 | 321 |
fun after_qed [[th]] = snd o typedef_result th; |
36323
655e2d74de3a
modernized naming conventions of main Isar proof elements;
wenzelm
parents:
36153
diff
changeset
|
322 |
in Proof.theorem NONE after_qed [[(goal, [goal_pat])]] lthy' end; |
17339 | 323 |
|
324 |
in |
|
6383 | 325 |
|
35840
01d7c4ba9050
allow sort constraints in HOL/typedef and related HOLCF variants;
wenzelm
parents:
35836
diff
changeset
|
326 |
val typedef = gen_typedef Syntax.check_term (K I); |
01d7c4ba9050
allow sort constraints in HOL/typedef and related HOLCF variants;
wenzelm
parents:
35836
diff
changeset
|
327 |
val typedef_cmd = gen_typedef Syntax.read_term Typedecl.read_constraint; |
17339 | 328 |
|
19705 | 329 |
end; |
15259 | 330 |
|
331 |
||
332 |
||
6383 | 333 |
(** outer syntax **) |
334 |
||
24867 | 335 |
val _ = |
67149 | 336 |
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:
46949
diff
changeset
|
337 |
"HOL type definition (requires non-emptiness proof)" |
61260 | 338 |
(Parse_Spec.overloaded -- Parse.type_args_constrained -- Parse.binding -- Parse.opt_mixfix -- |
67149 | 339 |
(\<^keyword>\<open>=\<close> |-- Parse.term) -- |
340 |
Scan.option (\<^keyword>\<open>morphisms\<close> |-- Parse.!!! (Parse.binding -- Parse.binding)) |
|
61260 | 341 |
>> (fn (((((overloaded, vs), t), mx), A), opt_morphs) => fn lthy => |
342 |
typedef_cmd {overloaded = overloaded} (t, vs, mx) A |
|
343 |
(SOME (make_morphisms t opt_morphs)) lthy)); |
|
344 |
||
345 |
||
346 |
val overloaded = typedef_overloaded; |
|
6357 | 347 |
|
68264 | 348 |
|
349 |
||
350 |
(** theory export **) |
|
351 |
||
352 |
val _ = Export_Theory.setup_presentation (fn _ => fn thy => |
|
353 |
let |
|
354 |
val parent_spaces = map Sign.type_space (Theory.parents_of thy); |
|
355 |
val typedefs = |
|
356 |
Name_Space.dest_table (#types (Type.rep_tsig (Sign.tsig_of thy))) |
|
357 |
|> maps (fn (name, _) => |
|
358 |
if exists (fn space => Name_Space.declared space name) parent_spaces then [] |
|
359 |
else |
|
360 |
get_info_global thy name |
|
361 |
|> map (fn ({rep_type, abs_type, Rep_name, Abs_name, axiom_name}, _) => |
|
362 |
(name, (rep_type, (abs_type, (Rep_name, (Abs_name, axiom_name))))))); |
|
363 |
in |
|
364 |
if null typedefs then () |
|
365 |
else |
|
366 |
Export_Theory.export_body thy "typedefs" |
|
367 |
let open XML.Encode Term_XML.Encode |
|
368 |
in list (pair string (pair typ (pair typ (pair string (pair string string))))) typedefs end |
|
369 |
end); |
|
370 |
||
29056 | 371 |
end; |