author | wenzelm |
Sat, 14 Jan 2012 17:45:04 +0100 | |
changeset 46215 | 0da9433f959e |
parent 45407 | a83574606719 |
child 46947 | b8c7eb0c2f89 |
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} * |
35994
9cc3df9a606e
Typedef.info: separate global and local part, only the latter is transformed by morphisms;
wenzelm
parents:
35840
diff
changeset
|
12 |
{inhabited: thm, type_definition: thm, set_def: thm option, Rep: thm, Rep_inverse: thm, |
29061
c67cc9402ba9
inhabitance goal is now stated in original form and result contracted --
wenzelm
parents:
29059
diff
changeset
|
13 |
Abs_inverse: thm, Rep_inject: thm, Abs_inject: thm, Rep_cases: thm, Abs_cases: thm, |
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 |
|
18 |
val interpretation: (string -> theory -> theory) -> theory -> theory |
|
19 |
val setup: theory -> theory |
|
35840
01d7c4ba9050
allow sort constraints in HOL/typedef and related HOLCF variants;
wenzelm
parents:
35836
diff
changeset
|
20 |
val add_typedef: bool -> binding option -> binding * (string * sort) list * mixfix -> |
35741 | 21 |
term -> (binding * binding) option -> tactic -> local_theory -> (string * info) * local_theory |
35840
01d7c4ba9050
allow sort constraints in HOL/typedef and related HOLCF variants;
wenzelm
parents:
35836
diff
changeset
|
22 |
val add_typedef_global: bool -> binding option -> binding * (string * sort) list * mixfix -> |
30345 | 23 |
term -> (binding * binding) option -> tactic -> theory -> (string * info) * theory |
35840
01d7c4ba9050
allow sort constraints in HOL/typedef and related HOLCF variants;
wenzelm
parents:
35836
diff
changeset
|
24 |
val typedef: (bool * binding) * (binding * (string * sort) list * mixfix) * term * |
35741 | 25 |
(binding * binding) option -> local_theory -> Proof.state |
35840
01d7c4ba9050
allow sort constraints in HOL/typedef and related HOLCF variants;
wenzelm
parents:
35836
diff
changeset
|
26 |
val typedef_cmd: (bool * binding) * (binding * (string * string option) list * mixfix) * string * |
35741 | 27 |
(binding * binding) option -> local_theory -> Proof.state |
4866 | 28 |
end; |
29 |
||
31723
f5cafe803b55
discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents:
30345
diff
changeset
|
30 |
structure Typedef: TYPEDEF = |
4866 | 31 |
struct |
32 |
||
17922 | 33 |
(** type definitions **) |
34 |
||
35 |
(* theory data *) |
|
15259 | 36 |
|
19705 | 37 |
type info = |
35994
9cc3df9a606e
Typedef.info: separate global and local part, only the latter is transformed by morphisms;
wenzelm
parents:
35840
diff
changeset
|
38 |
(*global part*) |
36107 | 39 |
{rep_type: typ, abs_type: typ, Rep_name: string, Abs_name: string, axiom_name: string} * |
35741 | 40 |
(*local part*) |
35994
9cc3df9a606e
Typedef.info: separate global and local part, only the latter is transformed by morphisms;
wenzelm
parents:
35840
diff
changeset
|
41 |
{inhabited: thm, type_definition: thm, set_def: thm option, Rep: thm, Rep_inverse: thm, |
9cc3df9a606e
Typedef.info: separate global and local part, only the latter is transformed by morphisms;
wenzelm
parents:
35840
diff
changeset
|
42 |
Abs_inverse: thm, Rep_inject: thm, Abs_inject: thm, Rep_cases: thm, Abs_cases: thm, |
9cc3df9a606e
Typedef.info: separate global and local part, only the latter is transformed by morphisms;
wenzelm
parents:
35840
diff
changeset
|
43 |
Rep_induct: thm, Abs_induct: thm}; |
19459 | 44 |
|
35741 | 45 |
fun transform_info phi (info: info) = |
46 |
let |
|
47 |
val thm = Morphism.thm phi; |
|
35994
9cc3df9a606e
Typedef.info: separate global and local part, only the latter is transformed by morphisms;
wenzelm
parents:
35840
diff
changeset
|
48 |
val (global_info, {inhabited, type_definition, |
35741 | 49 |
set_def, Rep, Rep_inverse, Abs_inverse, Rep_inject, Abs_inject, |
35994
9cc3df9a606e
Typedef.info: separate global and local part, only the latter is transformed by morphisms;
wenzelm
parents:
35840
diff
changeset
|
50 |
Rep_cases, Abs_cases, Rep_induct, Abs_induct}) = info; |
35741 | 51 |
in |
35994
9cc3df9a606e
Typedef.info: separate global and local part, only the latter is transformed by morphisms;
wenzelm
parents:
35840
diff
changeset
|
52 |
(global_info, |
9cc3df9a606e
Typedef.info: separate global and local part, only the latter is transformed by morphisms;
wenzelm
parents:
35840
diff
changeset
|
53 |
{inhabited = thm inhabited, type_definition = thm type_definition, |
9cc3df9a606e
Typedef.info: separate global and local part, only the latter is transformed by morphisms;
wenzelm
parents:
35840
diff
changeset
|
54 |
set_def = Option.map thm set_def, Rep = thm Rep, Rep_inverse = thm Rep_inverse, |
9cc3df9a606e
Typedef.info: separate global and local part, only the latter is transformed by morphisms;
wenzelm
parents:
35840
diff
changeset
|
55 |
Abs_inverse = thm Abs_inverse, Rep_inject = thm Rep_inject, Abs_inject = thm Abs_inject, |
9cc3df9a606e
Typedef.info: separate global and local part, only the latter is transformed by morphisms;
wenzelm
parents:
35840
diff
changeset
|
56 |
Rep_cases = thm Rep_cases, Abs_cases = thm Abs_cases, Rep_induct = thm Rep_induct, |
9cc3df9a606e
Typedef.info: separate global and local part, only the latter is transformed by morphisms;
wenzelm
parents:
35840
diff
changeset
|
57 |
Abs_induct = thm Abs_induct}) |
35741 | 58 |
end; |
59 |
||
60 |
structure Data = Generic_Data |
|
22846 | 61 |
( |
35741 | 62 |
type T = info list Symtab.table; |
15259 | 63 |
val empty = Symtab.empty; |
16458 | 64 |
val extend = I; |
35741 | 65 |
fun merge data = Symtab.merge_list (K true) data; |
22846 | 66 |
); |
15259 | 67 |
|
35741 | 68 |
val get_info = Symtab.lookup_list o Data.get o Context.Proof; |
69 |
val get_info_global = Symtab.lookup_list o Data.get o Context.Theory; |
|
70 |
||
71 |
fun put_info name info = Data.map (Symtab.cons_list (name, info)); |
|
72 |
||
73 |
||
74 |
(* global interpretation *) |
|
75 |
||
76 |
structure Typedef_Interpretation = Interpretation(type T = string val eq = op =); |
|
77 |
val interpretation = Typedef_Interpretation.interpretation; |
|
78 |
||
79 |
val setup = Typedef_Interpretation.init; |
|
80 |
||
81 |
||
82 |
(* primitive typedef axiomatization -- for fresh typedecl *) |
|
83 |
||
84 |
fun mk_inhabited A = |
|
85 |
let val T = HOLogic.dest_setT (Term.fastype_of A) |
|
86 |
in HOLogic.mk_Trueprop (HOLogic.exists_const T $ Abs ("x", T, HOLogic.mk_mem (Bound 0, A))) end; |
|
87 |
||
88 |
fun mk_typedef newT oldT RepC AbsC A = |
|
89 |
let |
|
90 |
val typedefC = |
|
91 |
Const (@{const_name type_definition}, |
|
92 |
(newT --> oldT) --> (oldT --> newT) --> HOLogic.mk_setT oldT --> HOLogic.boolT); |
|
93 |
in Logic.mk_implies (mk_inhabited A, HOLogic.mk_Trueprop (typedefC $ RepC $ AbsC $ A)) end; |
|
35134 | 94 |
|
42375
774df7c59508
report Name_Space.declare/define, relatively to context;
wenzelm
parents:
42361
diff
changeset
|
95 |
fun primitive_typedef typedef_name newT oldT Rep_name Abs_name A lthy = |
35741 | 96 |
let |
97 |
(* errors *) |
|
98 |
||
99 |
fun show_names pairs = commas_quote (map fst pairs); |
|
100 |
||
101 |
val lhs_tfrees = Term.add_tfreesT newT []; |
|
102 |
val rhs_tfrees = Term.add_tfreesT oldT []; |
|
103 |
val _ = |
|
104 |
(case fold (remove (op =)) lhs_tfrees rhs_tfrees of [] => () |
|
105 |
| extras => error ("Extra type variables in representing set: " ^ show_names extras)); |
|
106 |
||
107 |
val _ = |
|
108 |
(case Term.add_frees A [] of [] => [] |
|
109 |
| xs => error ("Illegal variables in representing set: " ^ show_names xs)); |
|
35134 | 110 |
|
35741 | 111 |
|
112 |
(* axiomatization *) |
|
113 |
||
42375
774df7c59508
report Name_Space.declare/define, relatively to context;
wenzelm
parents:
42361
diff
changeset
|
114 |
val ((RepC, AbsC), consts_lthy) = lthy |
774df7c59508
report Name_Space.declare/define, relatively to context;
wenzelm
parents:
42361
diff
changeset
|
115 |
|> Local_Theory.background_theory_result |
774df7c59508
report Name_Space.declare/define, relatively to context;
wenzelm
parents:
42361
diff
changeset
|
116 |
(Sign.declare_const lthy ((Rep_name, newT --> oldT), NoSyn) ##>> |
774df7c59508
report Name_Space.declare/define, relatively to context;
wenzelm
parents:
42361
diff
changeset
|
117 |
Sign.declare_const lthy ((Abs_name, oldT --> newT), NoSyn)); |
35741 | 118 |
|
119 |
val typedef_deps = Term.add_consts A []; |
|
120 |
||
42375
774df7c59508
report Name_Space.declare/define, relatively to context;
wenzelm
parents:
42361
diff
changeset
|
121 |
val ((axiom_name, axiom), axiom_lthy) = consts_lthy |
774df7c59508
report Name_Space.declare/define, relatively to context;
wenzelm
parents:
42361
diff
changeset
|
122 |
|> Local_Theory.background_theory_result |
774df7c59508
report Name_Space.declare/define, relatively to context;
wenzelm
parents:
42361
diff
changeset
|
123 |
(Thm.add_axiom consts_lthy (typedef_name, mk_typedef newT oldT RepC AbsC A) ##> |
774df7c59508
report Name_Space.declare/define, relatively to context;
wenzelm
parents:
42361
diff
changeset
|
124 |
Theory.add_deps consts_lthy "" (dest_Const RepC) typedef_deps ##> |
774df7c59508
report Name_Space.declare/define, relatively to context;
wenzelm
parents:
42361
diff
changeset
|
125 |
Theory.add_deps consts_lthy "" (dest_Const AbsC) typedef_deps); |
35741 | 126 |
|
42375
774df7c59508
report Name_Space.declare/define, relatively to context;
wenzelm
parents:
42361
diff
changeset
|
127 |
in ((RepC, AbsC, axiom_name, axiom), axiom_lthy) end; |
15259 | 128 |
|
129 |
||
6383 | 130 |
(* prepare_typedef *) |
131 |
||
35840
01d7c4ba9050
allow sort constraints in HOL/typedef and related HOLCF variants;
wenzelm
parents:
35836
diff
changeset
|
132 |
fun prepare_typedef prep_term def_set name (tname, raw_args, mx) raw_set opt_morphs lthy = |
4866 | 133 |
let |
35741 | 134 |
val full_name = Local_Theory.full_name lthy name; |
30345 | 135 |
val bname = Binding.name_of name; |
4866 | 136 |
|
35741 | 137 |
|
138 |
(* rhs *) |
|
139 |
||
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
|
140 |
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
|
141 |
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
|
142 |
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
|
143 |
|
21352 | 144 |
val setT = Term.fastype_of set; |
35741 | 145 |
val oldT = HOLogic.dest_setT setT handle TYPE _ => |
146 |
error ("Not a set type: " ^ quote (Syntax.string_of_typ lthy setT)); |
|
147 |
||
148 |
val goal = mk_inhabited set; |
|
42290
b1f544c84040
discontinued special treatment of structure Lexicon;
wenzelm
parents:
38757
diff
changeset
|
149 |
val goal_pat = mk_inhabited (Var (the_default (bname, 0) (Lexicon.read_variable bname), setT)); |
35741 | 150 |
|
151 |
||
152 |
(* lhs *) |
|
153 |
||
42361 | 154 |
val args = map (Proof_Context.check_tfree tmp_ctxt') raw_args; |
35741 | 155 |
val (newT, typedecl_lthy) = lthy |
35836
9380fab5f4f7
eliminated slightly odd typedecl_wrt in favour of explicit predeclare_constraints;
wenzelm
parents:
35766
diff
changeset
|
156 |
|> Typedecl.typedecl (tname, args, mx) |
35741 | 157 |
||> Variable.declare_term set; |
158 |
||
159 |
val Type (full_tname, type_args) = newT; |
|
160 |
val lhs_tfrees = map Term.dest_TFree type_args; |
|
161 |
||
162 |
||
163 |
(* set definition *) |
|
164 |
||
165 |
val ((set', set_def), set_lthy) = |
|
166 |
if def_set then |
|
167 |
typedecl_lthy |
|
35766 | 168 |
|> Local_Theory.define ((name, NoSyn), ((Thm.def_binding name, []), set)) |
169 |
|>> (fn (set', (_, set_def)) => (set', SOME set_def)) |
|
35741 | 170 |
else ((set, NONE), typedecl_lthy); |
171 |
||
172 |
||
173 |
(* axiomatization *) |
|
4866 | 174 |
|
30345 | 175 |
val (Rep_name, Abs_name) = |
176 |
(case opt_morphs of |
|
177 |
NONE => (Binding.prefix_name "Rep_" name, Binding.prefix_name "Abs_" name) |
|
178 |
| SOME morphs => morphs); |
|
10280 | 179 |
|
30345 | 180 |
val typedef_name = Binding.prefix_name "type_definition_" name; |
4866 | 181 |
|
36107 | 182 |
val ((RepC, AbsC, axiom_name, typedef), typedef_lthy) = |
35741 | 183 |
let |
42361 | 184 |
val thy = Proof_Context.theory_of set_lthy; |
35741 | 185 |
val cert = Thm.cterm_of thy; |
45407
a83574606719
more specific treatment of defines/assumes -- avoid normalizing defs by themselves (NB: locale specifications and Local_Theory.define may lead to arbitrary mixture);
wenzelm
parents:
45291
diff
changeset
|
186 |
val ((defs, _), A) = |
42361 | 187 |
Local_Defs.export_cterm set_lthy (Proof_Context.init_global thy) (cert set') |
36610
bafd82950e24
renamed ProofContext.init to ProofContext.init_global to emphasize that this is not the real thing;
wenzelm
parents:
36323
diff
changeset
|
188 |
||> Thm.term_of; |
18358 | 189 |
|
42375
774df7c59508
report Name_Space.declare/define, relatively to context;
wenzelm
parents:
42361
diff
changeset
|
190 |
val ((RepC, AbsC, axiom_name, axiom), axiom_lthy) = set_lthy |
774df7c59508
report Name_Space.declare/define, relatively to context;
wenzelm
parents:
42361
diff
changeset
|
191 |
|> primitive_typedef typedef_name newT oldT Rep_name Abs_name A; |
35741 | 192 |
|
42361 | 193 |
val cert = Thm.cterm_of (Proof_Context.theory_of axiom_lthy); |
35741 | 194 |
val typedef = |
195 |
Local_Defs.contract axiom_lthy defs (cert (mk_typedef newT oldT RepC AbsC set')) axiom; |
|
36107 | 196 |
in ((RepC, AbsC, axiom_name, typedef), axiom_lthy) end; |
35741 | 197 |
|
198 |
val alias_lthy = typedef_lthy |
|
199 |
|> Local_Theory.const_alias Rep_name (#1 (Term.dest_Const RepC)) |
|
200 |
|> Local_Theory.const_alias Abs_name (#1 (Term.dest_Const AbsC)); |
|
6383 | 201 |
|
29056 | 202 |
|
35741 | 203 |
(* result *) |
4866 | 204 |
|
35741 | 205 |
fun note_qualify ((b, atts), th) = |
206 |
Local_Theory.note ((Binding.qualify false bname b, map (Attrib.internal o K) atts), [th]) |
|
207 |
#>> (fn (_, [th']) => th'); |
|
4866 | 208 |
|
35741 | 209 |
fun typedef_result inhabited lthy1 = |
210 |
let |
|
42361 | 211 |
val cert = Thm.cterm_of (Proof_Context.theory_of lthy1); |
35741 | 212 |
val inhabited' = |
213 |
Local_Defs.contract lthy1 (the_list set_def) (cert (mk_inhabited set')) inhabited; |
|
214 |
val typedef' = inhabited' RS typedef; |
|
215 |
fun make th = Goal.norm_result (typedef' RS th); |
|
216 |
val (((((((((((_, [type_definition]), Rep), Rep_inverse), Abs_inverse), Rep_inject), |
|
217 |
Abs_inject), Rep_cases), Abs_cases), Rep_induct), Abs_induct), lthy2) = lthy1 |
|
218 |
|> Local_Theory.note ((typedef_name, []), [typedef']) |
|
219 |
||>> note_qualify ((Rep_name, []), make @{thm type_definition.Rep}) |
|
220 |
||>> note_qualify ((Binding.suffix_name "_inverse" Rep_name, []), |
|
221 |
make @{thm type_definition.Rep_inverse}) |
|
222 |
||>> note_qualify ((Binding.suffix_name "_inverse" Abs_name, []), |
|
223 |
make @{thm type_definition.Abs_inverse}) |
|
224 |
||>> note_qualify ((Binding.suffix_name "_inject" Rep_name, []), |
|
225 |
make @{thm type_definition.Rep_inject}) |
|
226 |
||>> note_qualify ((Binding.suffix_name "_inject" Abs_name, []), |
|
227 |
make @{thm type_definition.Abs_inject}) |
|
228 |
||>> note_qualify ((Binding.suffix_name "_cases" Rep_name, |
|
229 |
[Rule_Cases.case_names [Binding.name_of Rep_name], Induct.cases_pred full_name]), |
|
230 |
make @{thm type_definition.Rep_cases}) |
|
231 |
||>> note_qualify ((Binding.suffix_name "_cases" Abs_name, |
|
232 |
[Rule_Cases.case_names [Binding.name_of Abs_name], Induct.cases_type full_tname]), |
|
233 |
make @{thm type_definition.Abs_cases}) |
|
234 |
||>> note_qualify ((Binding.suffix_name "_induct" Rep_name, |
|
235 |
[Rule_Cases.case_names [Binding.name_of Rep_name], Induct.induct_pred full_name]), |
|
236 |
make @{thm type_definition.Rep_induct}) |
|
237 |
||>> note_qualify ((Binding.suffix_name "_induct" Abs_name, |
|
238 |
[Rule_Cases.case_names [Binding.name_of Abs_name], Induct.induct_type full_tname]), |
|
239 |
make @{thm type_definition.Abs_induct}); |
|
4866 | 240 |
|
35994
9cc3df9a606e
Typedef.info: separate global and local part, only the latter is transformed by morphisms;
wenzelm
parents:
35840
diff
changeset
|
241 |
val info = |
36107 | 242 |
({rep_type = oldT, abs_type = newT, Rep_name = #1 (Term.dest_Const RepC), |
243 |
Abs_name = #1 (Term.dest_Const AbsC), axiom_name = axiom_name}, |
|
35994
9cc3df9a606e
Typedef.info: separate global and local part, only the latter is transformed by morphisms;
wenzelm
parents:
35840
diff
changeset
|
244 |
{inhabited = inhabited, type_definition = type_definition, set_def = set_def, |
35741 | 245 |
Rep = Rep, Rep_inverse = Rep_inverse, Abs_inverse = Abs_inverse, |
246 |
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
|
247 |
Abs_cases = Abs_cases, Rep_induct = Rep_induct, Abs_induct = Abs_induct}); |
35741 | 248 |
in |
249 |
lthy2 |
|
45291
57cd50f98fdc
uniform Local_Theory.declaration with explicit params;
wenzelm
parents:
42381
diff
changeset
|
250 |
|> Local_Theory.declaration {syntax = false, pervasive = true} |
57cd50f98fdc
uniform Local_Theory.declaration with explicit params;
wenzelm
parents:
42381
diff
changeset
|
251 |
(fn phi => put_info full_tname (transform_info phi info)) |
38757
2b3e054ae6fc
renamed Local_Theory.theory(_result) to Local_Theory.background_theory(_result) to emphasize that this belongs to the infrastructure and is rarely appropriate in user-space tools;
wenzelm
parents:
38388
diff
changeset
|
252 |
|> Local_Theory.background_theory (Typedef_Interpretation.data full_tname) |
35741 | 253 |
|> pair (full_tname, info) |
254 |
end; |
|
11426
f280d4b29a2c
abtract non-emptiness statements (no longer use Eps);
wenzelm
parents:
10697
diff
changeset
|
255 |
|
35741 | 256 |
in ((goal, goal_pat, typedef_result), alias_lthy) end |
30345 | 257 |
handle ERROR msg => |
42381
309ec68442c6
added Binding.print convenience, which includes quote already;
wenzelm
parents:
42375
diff
changeset
|
258 |
cat_error msg ("The error(s) above occurred in typedef " ^ Binding.print name); |
4866 | 259 |
|
260 |
||
29056 | 261 |
(* add_typedef: tactic interface *) |
4866 | 262 |
|
35741 | 263 |
fun add_typedef def opt_name typ set opt_morphs tac lthy = |
6383 | 264 |
let |
17922 | 265 |
val name = the_default (#1 typ) opt_name; |
35741 | 266 |
val ((goal, _, typedef_result), lthy') = |
267 |
prepare_typedef Syntax.check_term def name typ set opt_morphs lthy; |
|
268 |
val inhabited = |
|
269 |
Goal.prove lthy' [] [] goal (K tac) |
|
270 |
|> Goal.norm_result |> Thm.close_derivation; |
|
271 |
in typedef_result inhabited lthy' end; |
|
272 |
||
273 |
fun add_typedef_global def opt_name typ set opt_morphs tac = |
|
38388 | 274 |
Named_Target.theory_init |
35741 | 275 |
#> add_typedef def opt_name typ set opt_morphs tac |
276 |
#> Local_Theory.exit_result_global (apsnd o transform_info); |
|
4866 | 277 |
|
17339 | 278 |
|
29056 | 279 |
(* typedef: proof interface *) |
6383 | 280 |
|
17339 | 281 |
local |
282 |
||
35840
01d7c4ba9050
allow sort constraints in HOL/typedef and related HOLCF variants;
wenzelm
parents:
35836
diff
changeset
|
283 |
fun gen_typedef prep_term prep_constraint ((def, name), (b, raw_args, mx), set, opt_morphs) lthy = |
11822 | 284 |
let |
35840
01d7c4ba9050
allow sort constraints in HOL/typedef and related HOLCF variants;
wenzelm
parents:
35836
diff
changeset
|
285 |
val args = map (apsnd (prep_constraint lthy)) raw_args; |
35741 | 286 |
val ((goal, goal_pat, typedef_result), lthy') = |
35840
01d7c4ba9050
allow sort constraints in HOL/typedef and related HOLCF variants;
wenzelm
parents:
35836
diff
changeset
|
287 |
prepare_typedef prep_term def name (b, args, mx) set opt_morphs lthy; |
35741 | 288 |
fun after_qed [[th]] = snd o typedef_result th; |
36323
655e2d74de3a
modernized naming conventions of main Isar proof elements;
wenzelm
parents:
36153
diff
changeset
|
289 |
in Proof.theorem NONE after_qed [[(goal, [goal_pat])]] lthy' end; |
17339 | 290 |
|
291 |
in |
|
6383 | 292 |
|
35840
01d7c4ba9050
allow sort constraints in HOL/typedef and related HOLCF variants;
wenzelm
parents:
35836
diff
changeset
|
293 |
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
|
294 |
val typedef_cmd = gen_typedef Syntax.read_term Typedecl.read_constraint; |
17339 | 295 |
|
19705 | 296 |
end; |
15259 | 297 |
|
298 |
||
299 |
||
6383 | 300 |
(** outer syntax **) |
301 |
||
36960
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents:
36610
diff
changeset
|
302 |
val _ = Keyword.keyword "morphisms"; |
24867 | 303 |
|
304 |
val _ = |
|
36960
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents:
36610
diff
changeset
|
305 |
Outer_Syntax.local_theory_to_proof "typedef" "HOL type definition (requires non-emptiness proof)" |
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents:
36610
diff
changeset
|
306 |
Keyword.thy_goal |
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents:
36610
diff
changeset
|
307 |
(Scan.optional (Parse.$$$ "(" |-- |
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents:
36610
diff
changeset
|
308 |
((Parse.$$$ "open" >> K false) -- Scan.option Parse.binding || |
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents:
36610
diff
changeset
|
309 |
Parse.binding >> (fn s => (true, SOME s))) --| Parse.$$$ ")") (true, NONE) -- |
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents:
36610
diff
changeset
|
310 |
(Parse.type_args_constrained -- Parse.binding) -- |
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents:
36610
diff
changeset
|
311 |
Parse.opt_mixfix -- (Parse.$$$ "=" |-- Parse.term) -- |
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents:
36610
diff
changeset
|
312 |
Scan.option (Parse.$$$ "morphisms" |-- Parse.!!! (Parse.binding -- Parse.binding)) |
35840
01d7c4ba9050
allow sort constraints in HOL/typedef and related HOLCF variants;
wenzelm
parents:
35836
diff
changeset
|
313 |
>> (fn ((((((def, opt_name), (args, t)), mx), A), morphs)) => |
01d7c4ba9050
allow sort constraints in HOL/typedef and related HOLCF variants;
wenzelm
parents:
35836
diff
changeset
|
314 |
typedef_cmd ((def, the_default t opt_name), (t, args, mx), A, morphs))); |
6357 | 315 |
|
29056 | 316 |
end; |
317 |