1 (* Title: HOL/Tools/typedef_package.ML
3 Author: Markus Wenzel and Stefan Berghofer, TU Muenchen
5 Gordon/HOL-style type definitions.
8 signature TYPEDEF_PACKAGE =
10 val quiet_mode: bool ref
11 val add_typedecls: (bstring * string list * mixfix) list -> theory -> theory
12 val add_typedef_x: string -> bstring * string list * mixfix ->
13 string -> string list -> thm list -> tactic option -> theory -> theory
14 val add_typedef: bool -> string option -> bstring * string list * mixfix ->
15 string -> (bstring * bstring) option -> tactic -> theory -> theory *
16 {type_definition: thm, set_def: thm option, Rep: thm, Rep_inverse: thm,
17 Abs_inverse: thm, Rep_inject: thm, Abs_inject: thm, Rep_cases: thm, Abs_cases: thm,
18 Rep_induct: thm, Abs_induct: thm}
19 val add_typedef_i: bool -> string option -> bstring * string list * mixfix ->
20 term -> (bstring * bstring) option -> tactic -> theory -> theory *
21 {type_definition: thm, set_def: thm option, Rep: thm, Rep_inverse: thm,
22 Abs_inverse: thm, Rep_inject: thm, Abs_inject: thm, Rep_cases: thm, Abs_cases: thm,
23 Rep_induct: thm, Abs_induct: thm}
24 val typedef_proof: (bool * string) * (bstring * string list * mixfix) * string
25 * (string * string) option -> bool -> theory -> ProofHistory.T
26 val typedef_proof_i: (bool * string) * (bstring * string list * mixfix) * term
27 * (string * string) option -> bool -> theory -> ProofHistory.T
28 val setup: (theory -> theory) list
31 structure TypedefPackage: TYPEDEF_PACKAGE =
35 (** theory context references **)
37 val type_definitionN = "Typedef.type_definition";
39 val Rep = thm "type_definition.Rep";
40 val Rep_inverse = thm "type_definition.Rep_inverse";
41 val Abs_inverse = thm "type_definition.Abs_inverse";
42 val Rep_inject = thm "type_definition.Rep_inject";
43 val Abs_inject = thm "type_definition.Abs_inject";
44 val Rep_cases = thm "type_definition.Rep_cases";
45 val Abs_cases = thm "type_definition.Abs_cases";
46 val Rep_induct = thm "type_definition.Rep_induct";
47 val Abs_induct = thm "type_definition.Abs_induct";
53 structure TypedefData = TheoryDataFun
55 val name = "HOL/typedef";
56 type T = (typ * typ * string * string) Symtab.table;
57 val empty = Symtab.empty;
60 fun merge _ (tabs: T * T) = Symtab.merge (op =) tabs;
64 fun put_typedef newT oldT Abs_name Rep_name thy =
65 TypedefData.put (Symtab.update_new
66 ((fst (dest_Type newT), (newT, oldT, Abs_name, Rep_name)),
67 TypedefData.get thy)) thy;
71 (** type declarations **)
73 fun add_typedecls decls thy =
75 fun arity_of (raw_name, args, mx) =
76 (Sign.full_name thy (Syntax.type_name raw_name mx),
77 replicate (length args) HOLogic.typeS, HOLogic.typeS);
79 if can (Theory.assert_super HOL.thy) thy then
80 thy |> Theory.add_typedecls decls
81 |> Theory.add_arities_i (map arity_of decls)
82 else thy |> Theory.add_typedecls decls
87 (** type definitions **)
91 val quiet_mode = ref false;
92 fun message s = if ! quiet_mode then () else writeln s;
95 (* prove_nonempty -- tactical version *) (*exception ERROR*)
97 fun prove_nonempty thy cset goal (witn1_tac, witn_names, witn_thms, witn2_tac) =
99 val is_def = Logic.is_equals o #prop o Thm.rep_thm;
100 val thms = PureThy.get_thmss thy (map Name witn_names) @ witn_thms;
103 TRY (rewrite_goals_tac (List.filter is_def thms)) THEN
104 TRY (REPEAT_FIRST (resolve_tac (filter_out is_def thms))) THEN
105 getOpt (witn2_tac, TRY (ALLGOALS (CLASET' blast_tac)));
107 message ("Proving non-emptiness of set " ^ quote (string_of_cterm cset) ^ " ...");
108 Tactic.prove thy [] [] goal (K tac)
110 handle ERROR => error ("Failed to prove non-emptiness of " ^ quote (string_of_cterm cset));
113 (* prepare_typedef *)
115 fun read_term thy used s =
116 #1 (Thm.read_def_cterm (thy, K NONE, K NONE) used true (s, HOLogic.typeT));
118 fun cert_term thy _ t = Thm.cterm_of thy t handle TERM (msg, _) => error msg;
120 fun err_in_typedef name =
121 error ("The error(s) above occurred in typedef " ^ quote name);
123 fun prepare_typedef prep_term def name (t, vs, mx) raw_set opt_morphs thy =
125 val _ = Theory.requires thy "Typedef" "typedefs";
126 val full = Sign.full_name thy;
129 val full_name = full name;
130 val cset = prep_term thy vs raw_set;
131 val {T = setT, t = set, ...} = Thm.rep_cterm cset;
132 val rhs_tfrees = term_tfrees set;
133 val oldT = HOLogic.dest_setT setT handle TYPE _ =>
134 error ("Not a set type: " ^ quote (Sign.string_of_typ thy setT));
136 HOLogic.mk_Trueprop (HOLogic.mk_exists ("x", oldT, HOLogic.mk_mem (Free ("x", oldT), A)));
137 val goal = mk_nonempty set;
138 val vname = take_suffix Symbol.is_digit (Symbol.explode name)
139 |> apfst implode |> apsnd (#1 o Library.read_int);
140 val goal_pat = mk_nonempty (Var (vname, setT));
143 val lhs_tfrees = map (fn v => (v, getOpt (assoc (rhs_tfrees, v), HOLogic.typeS))) vs;
144 val tname = Syntax.type_name t mx;
145 val full_tname = full tname;
146 val newT = Type (full_tname, map TFree lhs_tfrees);
148 val (Rep_name, Abs_name) = getOpt (opt_morphs, ("Rep_" ^ name, "Abs_" ^ name));
149 val setC = Const (full_name, setT);
150 val RepC = Const (full Rep_name, newT --> oldT);
151 val AbsC = Const (full Abs_name, oldT --> newT);
152 val x_new = Free ("x", newT);
153 val y_old = Free ("y", oldT);
155 val set' = if def then setC else set;
157 val typedef_name = "type_definition_" ^ name;
159 Const (type_definitionN, (newT --> oldT) --> (oldT --> newT) --> setT --> HOLogic.boolT);
161 Logic.mk_implies (goal, HOLogic.mk_Trueprop (typedefC $ RepC $ AbsC $ set'));
163 fun typedef_result (theory, nonempty) =
165 |> put_typedef newT oldT (full Abs_name) (full Rep_name)
166 |> add_typedecls [(t, vs, mx)]
167 |> Theory.add_consts_i
168 ((if def then [(name, setT, NoSyn)] else []) @
169 [(Rep_name, newT --> oldT, NoSyn),
170 (Abs_name, oldT --> newT, NoSyn)])
171 |> (if def then (apsnd (SOME o hd) oo (PureThy.add_defs_i false o map Thm.no_attributes))
172 [Logic.mk_defpair (setC, set)]
174 |>>> PureThy.add_axioms_i [((typedef_name, typedef_prop),
175 [apsnd (fn cond_axm => Drule.standard (nonempty RS cond_axm))])]
176 |>> Theory.add_finals_i false [RepC, AbsC]
177 |> (fn (theory', (set_def, [type_definition])) =>
179 fun make th = Drule.standard (th OF [type_definition]);
180 val (theory'', [Rep, Rep_inverse, Abs_inverse, Rep_inject, Abs_inject,
181 Rep_cases, Abs_cases, Rep_induct, Abs_induct]) =
183 |> Theory.add_path name
185 ([((Rep_name, make Rep), []),
186 ((Rep_name ^ "_inverse", make Rep_inverse), []),
187 ((Abs_name ^ "_inverse", make Abs_inverse), []),
188 ((Rep_name ^ "_inject", make Rep_inject), []),
189 ((Abs_name ^ "_inject", make Abs_inject), []),
190 ((Rep_name ^ "_cases", make Rep_cases),
191 [RuleCases.case_names [Rep_name], InductAttrib.cases_set_global full_name]),
192 ((Abs_name ^ "_cases", make Abs_cases),
193 [RuleCases.case_names [Abs_name], InductAttrib.cases_type_global full_tname]),
194 ((Rep_name ^ "_induct", make Rep_induct),
195 [RuleCases.case_names [Rep_name], InductAttrib.induct_set_global full_name]),
196 ((Abs_name ^ "_induct", make Abs_induct),
197 [RuleCases.case_names [Abs_name], InductAttrib.induct_type_global full_tname])])
198 |>> Theory.parent_path;
199 val result = {type_definition = type_definition, set_def = set_def,
200 Rep = Rep, Rep_inverse = Rep_inverse, Abs_inverse = Abs_inverse,
201 Rep_inject = Rep_inject, Abs_inject = Abs_inject, Rep_cases = Rep_cases,
202 Abs_cases = Abs_cases, Rep_induct = Rep_induct, Abs_induct = Abs_induct};
203 in ((theory'', type_definition), result) end);
208 fun show_names pairs = commas_quote (map fst pairs);
211 if null (term_vars set) andalso null (term_tvars set) then []
212 else ["Illegal schematic variable(s) on rhs"];
215 (case duplicates lhs_tfrees of [] => []
216 | dups => ["Duplicate type variables on lhs: " ^ show_names dups]);
218 val extra_rhs_tfrees =
219 (case gen_rems (op =) (rhs_tfrees, lhs_tfrees) of [] => []
220 | extras => ["Extra type variables on rhs: " ^ show_names extras]);
223 (case term_frees set of [] => []
224 | xs => ["Illegal variables on rhs: " ^ show_names (map dest_Free xs)]);
226 val errs = illegal_vars @ dup_lhs_tfrees @ extra_rhs_tfrees @ illegal_frees;
227 val _ = if null errs then () else error (cat_lines errs);
229 (*test theory errors now!*)
230 val test_thy = Theory.copy thy;
232 setmp quick_and_dirty true (SkipProof.make_thm test_thy) goal) |> typedef_result;
234 in (cset, goal, goal_pat, typedef_result) end
235 handle ERROR => err_in_typedef name;
238 (* add_typedef interfaces *)
240 fun gen_typedef prep_term def name typ set opt_morphs tac1 names thms tac2 thy =
242 val (cset, goal, _, typedef_result) =
243 prepare_typedef prep_term def name typ set opt_morphs thy;
244 val non_empty = prove_nonempty thy cset goal (tac1, names, thms, tac2);
245 val ((thy', _), result) = (thy, non_empty) |> typedef_result;
246 in (thy', result) end;
248 fun sane_typedef prep_term def opt_name typ set opt_morphs tac =
249 gen_typedef prep_term def
250 (getOpt (opt_name, #1 typ)) typ set opt_morphs all_tac [] [] (SOME tac);
252 fun add_typedef_x name typ set names thms tac =
253 #1 o gen_typedef read_term true name typ set NONE (Tactic.rtac exI 1) names thms tac;
255 val add_typedef = sane_typedef read_term;
256 val add_typedef_i = sane_typedef cert_term;
259 (* typedef_proof interface *)
261 fun gen_typedef_proof prep_term ((def, name), typ, set, opt_morphs) int thy =
263 val (_, goal, goal_pat, att_result) =
264 prepare_typedef prep_term def name typ set opt_morphs thy;
265 val att = #1 o att_result;
266 in thy |> IsarThy.theorem_i Drule.internalK (("", [att]), (goal, ([goal_pat], []))) int end;
268 val typedef_proof = gen_typedef_proof read_term;
269 val typedef_proof_i = gen_typedef_proof cert_term;
273 (** trivial code generator **)
275 fun typedef_codegen thy gr dep brack t =
279 val (gr', _) = Codegen.invoke_tycodegen thy dep false (gr, T);
281 foldl_map (Codegen.invoke_codegen thy dep true) (gr', ts);
282 val id = Codegen.mk_const_id thy s
283 in SOME (gr'', Codegen.mk_app brack (Pretty.str id) ps) end;
284 fun get_name (Type (tname, _)) = tname
286 fun lookup f T = getOpt (Option.map f (Symtab.lookup
287 (TypedefData.get thy, get_name T)), "")
289 (case strip_comb t of
290 (Const (s, Type ("fun", [T, U])), ts) =>
291 if lookup #4 T = s andalso
292 is_none (Codegen.get_assoc_type thy (get_name T))
294 else if lookup #3 U = s andalso
295 is_none (Codegen.get_assoc_type thy (get_name U))
301 fun mk_tyexpr [] s = Pretty.str s
302 | mk_tyexpr [p] s = Pretty.block [p, Pretty.str (" " ^ s)]
303 | mk_tyexpr ps s = Pretty.list "(" (") " ^ s) ps;
305 fun typedef_tycodegen thy gr dep brack (Type (s, Ts)) =
306 (case Symtab.lookup (TypedefData.get thy, s) of
308 | SOME (newT as Type (tname, Us), oldT, Abs_name, Rep_name) =>
309 if isSome (Codegen.get_assoc_type thy tname) then NONE else
311 val Abs_id = Codegen.mk_const_id thy Abs_name;
312 val Rep_id = Codegen.mk_const_id thy Rep_name;
313 val ty_id = Codegen.mk_type_id thy s;
314 val (gr', qs) = foldl_map
315 (Codegen.invoke_tycodegen thy dep (length Ts = 1)) (gr, Ts);
316 val gr'' = Graph.add_edge (Abs_id, dep) gr' handle Graph.UNDEF _ =>
318 val (gr'', p :: ps) = foldl_map
319 (Codegen.invoke_tycodegen thy Abs_id false)
320 (Graph.add_edge (Abs_id, dep)
321 (Graph.new_node (Abs_id, (NONE, "")) gr'), oldT :: Us);
323 Pretty.string_of (Pretty.block [Pretty.str "datatype ",
325 Pretty.str " =", Pretty.brk 1, Pretty.str (Abs_id ^ " of"),
326 Pretty.brk 1, p, Pretty.str ";"]) ^ "\n\n" ^
327 Pretty.string_of (Pretty.block [Pretty.str ("fun " ^ Rep_id),
328 Pretty.brk 1, Pretty.str ("(" ^ Abs_id), Pretty.brk 1,
329 Pretty.str "x) = x;"]) ^ "\n\n" ^
330 (if "term_of" mem !Codegen.mode then
331 Pretty.string_of (Pretty.block [Pretty.str "fun ",
332 Codegen.mk_term_of thy false newT, Pretty.brk 1,
333 Pretty.str ("(" ^ Abs_id), Pretty.brk 1,
334 Pretty.str "x) =", Pretty.brk 1,
335 Pretty.block [Pretty.str ("Const (\"" ^ Abs_name ^ "\","),
336 Pretty.brk 1, Codegen.mk_type false (oldT --> newT),
337 Pretty.str ")"], Pretty.str " $", Pretty.brk 1,
338 Codegen.mk_term_of thy false oldT, Pretty.brk 1,
339 Pretty.str "x;"]) ^ "\n\n"
341 (if "test" mem !Codegen.mode then
342 Pretty.string_of (Pretty.block [Pretty.str "fun ",
343 Codegen.mk_gen thy false [] "" newT, Pretty.brk 1,
344 Pretty.str "i =", Pretty.brk 1,
345 Pretty.block [Pretty.str (Abs_id ^ " ("),
346 Codegen.mk_gen thy false [] "" oldT, Pretty.brk 1,
347 Pretty.str "i);"]]) ^ "\n\n"
349 in Graph.map_node Abs_id (K (NONE, s)) gr'' end
351 SOME (gr'', mk_tyexpr qs ty_id)
353 | typedef_tycodegen thy gr dep brack _ = NONE;
357 Codegen.add_codegen "typedef" typedef_codegen,
358 Codegen.add_tycodegen "typedef" typedef_tycodegen];
364 local structure P = OuterParse and K = OuterSyntax.Keyword in
367 OuterSyntax.command "typedecl" "type declaration (HOL)" K.thy_decl
368 (P.type_args -- P.name -- P.opt_infix >> (fn ((vs, t), mx) =>
369 Toplevel.theory (add_typedecls [(t, vs, mx)])));
372 val typedef_proof_decl =
373 Scan.optional (P.$$$ "(" |--
374 ((P.$$$ "open" >> K false) -- Scan.option P.name || P.name >> (fn s => (true, SOME s)))
375 --| P.$$$ ")") (true, NONE) --
376 (P.type_args -- P.name) -- P.opt_infix -- (P.$$$ "=" |-- P.term) --
377 Scan.option (P.$$$ "morphisms" |-- P.!!! (P.name -- P.name));
379 fun mk_typedef_proof ((((((def, opt_name), (vs, t)), mx), A), morphs)) =
380 typedef_proof ((def, getOpt (opt_name, Syntax.type_name t mx)), (t, vs, mx), A, morphs);
383 OuterSyntax.command "typedef" "HOL type definition (requires non-emptiness proof)" K.thy_goal
384 (typedef_proof_decl >> (Toplevel.print oo (Toplevel.theory_to_proof o mk_typedef_proof)));
387 val _ = OuterSyntax.add_keywords ["morphisms"];
388 val _ = OuterSyntax.add_parsers [typedeclP, typedefP];