wenzelm@4866
|
1 |
(* Title: HOL/Tools/typedef_package.ML
|
wenzelm@4866
|
2 |
ID: $Id$
|
wenzelm@4866
|
3 |
Author: Markus Wenzel, TU Muenchen
|
wenzelm@9230
|
4 |
License: GPL (GNU GENERAL PUBLIC LICENSE)
|
wenzelm@4866
|
5 |
|
wenzelm@6357
|
6 |
Gordon/HOL-style type definitions.
|
wenzelm@4866
|
7 |
*)
|
wenzelm@4866
|
8 |
|
wenzelm@4866
|
9 |
signature TYPEDEF_PACKAGE =
|
wenzelm@4866
|
10 |
sig
|
wenzelm@5697
|
11 |
val quiet_mode: bool ref
|
wenzelm@5104
|
12 |
val add_typedecls: (bstring * string list * mixfix) list -> theory -> theory
|
wenzelm@4866
|
13 |
val add_typedef: string -> bstring * string list * mixfix ->
|
wenzelm@4866
|
14 |
string -> string list -> thm list -> tactic option -> theory -> theory
|
wenzelm@4866
|
15 |
val add_typedef_i: string -> bstring * string list * mixfix ->
|
wenzelm@4866
|
16 |
term -> string list -> thm list -> tactic option -> theory -> theory
|
berghofe@5180
|
17 |
val add_typedef_i_no_def: string -> bstring * string list * mixfix ->
|
berghofe@5180
|
18 |
term -> string list -> thm list -> tactic option -> theory -> theory
|
wenzelm@6729
|
19 |
val typedef_proof: (string * (bstring * string list * mixfix) * string) * Comment.text
|
wenzelm@6690
|
20 |
-> bool -> theory -> ProofHistory.T
|
wenzelm@6729
|
21 |
val typedef_proof_i: (string * (bstring * string list * mixfix) * term) * Comment.text
|
wenzelm@6690
|
22 |
-> bool -> theory -> ProofHistory.T
|
wenzelm@4866
|
23 |
end;
|
wenzelm@4866
|
24 |
|
wenzelm@4866
|
25 |
structure TypedefPackage: TYPEDEF_PACKAGE =
|
wenzelm@4866
|
26 |
struct
|
wenzelm@4866
|
27 |
|
wenzelm@4866
|
28 |
|
wenzelm@5104
|
29 |
(** type declarations **)
|
wenzelm@5104
|
30 |
|
wenzelm@5104
|
31 |
fun add_typedecls decls thy =
|
wenzelm@5104
|
32 |
let
|
wenzelm@5104
|
33 |
val full = Sign.full_name (Theory.sign_of thy);
|
wenzelm@5104
|
34 |
|
wenzelm@5104
|
35 |
fun arity_of (raw_name, args, mx) =
|
wenzelm@5104
|
36 |
(full (Syntax.type_name raw_name mx), replicate (length args) HOLogic.termS, HOLogic.termS);
|
wenzelm@5104
|
37 |
in
|
wenzelm@8141
|
38 |
if can (Theory.assert_super HOL.thy) thy then
|
wenzelm@8141
|
39 |
thy
|
wenzelm@8141
|
40 |
|> PureThy.add_typedecls decls
|
wenzelm@8141
|
41 |
|> Theory.add_arities_i (map arity_of decls)
|
wenzelm@8141
|
42 |
else thy |> PureThy.add_typedecls decls
|
wenzelm@5104
|
43 |
end;
|
wenzelm@5104
|
44 |
|
wenzelm@5104
|
45 |
|
wenzelm@5104
|
46 |
|
wenzelm@5104
|
47 |
(** type definitions **)
|
wenzelm@5104
|
48 |
|
wenzelm@5697
|
49 |
(* messages *)
|
wenzelm@5697
|
50 |
|
wenzelm@5697
|
51 |
val quiet_mode = ref false;
|
wenzelm@5697
|
52 |
fun message s = if ! quiet_mode then () else writeln s;
|
wenzelm@5697
|
53 |
|
wenzelm@5697
|
54 |
|
wenzelm@6383
|
55 |
(* non-emptiness of set *) (*exception ERROR*)
|
wenzelm@4866
|
56 |
|
wenzelm@6383
|
57 |
fun check_nonempty cset thm =
|
wenzelm@4866
|
58 |
let
|
wenzelm@6383
|
59 |
val {t, sign, maxidx, ...} = Thm.rep_cterm cset;
|
wenzelm@6383
|
60 |
val {prop, ...} = Thm.rep_thm (Thm.transfer_sg sign (Drule.standard thm));
|
wenzelm@6383
|
61 |
val matches = Pattern.matches (Sign.tsig_of sign);
|
wenzelm@6383
|
62 |
in
|
wenzelm@6383
|
63 |
(case try (HOLogic.dest_mem o HOLogic.dest_Trueprop) prop of
|
wenzelm@6383
|
64 |
None => raise ERROR
|
wenzelm@6383
|
65 |
| Some (_, A) => if matches (Logic.incr_indexes ([], maxidx) A, t) then () else raise ERROR)
|
wenzelm@6383
|
66 |
end handle ERROR => error ("Bad non-emptiness theorem " ^ Display.string_of_thm thm ^
|
wenzelm@6383
|
67 |
"\nfor set " ^ quote (Display.string_of_cterm cset));
|
wenzelm@6383
|
68 |
|
wenzelm@7481
|
69 |
fun goal_nonempty ex cset =
|
wenzelm@6383
|
70 |
let
|
wenzelm@6383
|
71 |
val {T = setT, t = A, maxidx, sign} = Thm.rep_cterm cset;
|
wenzelm@4866
|
72 |
val T = HOLogic.dest_setT setT;
|
wenzelm@7481
|
73 |
val tm =
|
wenzelm@7481
|
74 |
if ex then HOLogic.mk_exists ("x", T, HOLogic.mk_mem (Free ("x", T), A))
|
wenzelm@7481
|
75 |
else HOLogic.mk_mem (Var (("x", maxidx + 1), T), A); (*old-style version*)
|
wenzelm@7481
|
76 |
in Thm.cterm_of sign (HOLogic.mk_Trueprop tm) end;
|
wenzelm@6383
|
77 |
|
wenzelm@6383
|
78 |
fun prove_nonempty thy cset (witn_names, witn_thms, witn_tac) =
|
wenzelm@6383
|
79 |
let
|
wenzelm@6383
|
80 |
val is_def = Logic.is_equals o #prop o Thm.rep_thm;
|
wenzelm@6383
|
81 |
val thms = PureThy.get_thmss thy witn_names @ witn_thms;
|
wenzelm@4866
|
82 |
val tac =
|
wenzelm@4866
|
83 |
TRY (rewrite_goals_tac (filter is_def thms)) THEN
|
wenzelm@4866
|
84 |
TRY (REPEAT_FIRST (resolve_tac (filter_out is_def thms))) THEN
|
wenzelm@6383
|
85 |
if_none witn_tac (TRY (ALLGOALS (CLASET' blast_tac)));
|
wenzelm@4866
|
86 |
in
|
wenzelm@6383
|
87 |
message ("Proving non-emptiness of set " ^ quote (string_of_cterm cset) ^ " ...");
|
wenzelm@7481
|
88 |
prove_goalw_cterm [] (goal_nonempty false cset) (K [tac])
|
wenzelm@6383
|
89 |
end handle ERROR => error ("Failed to prove non-emptiness of " ^ quote (string_of_cterm cset));
|
wenzelm@4866
|
90 |
|
wenzelm@4866
|
91 |
|
wenzelm@6383
|
92 |
(* prepare_typedef *)
|
wenzelm@6383
|
93 |
|
wenzelm@6383
|
94 |
fun read_term sg used s =
|
wenzelm@8100
|
95 |
#1 (Thm.read_def_cterm (sg, K None, K None) used true (s, HOLogic.termT));
|
wenzelm@4866
|
96 |
|
wenzelm@6383
|
97 |
fun cert_term sg _ t = Thm.cterm_of sg t handle TERM (msg, _) => error msg;
|
wenzelm@6383
|
98 |
|
wenzelm@6383
|
99 |
fun err_in_typedef name =
|
wenzelm@6383
|
100 |
error ("The error(s) above occurred in typedef " ^ quote name);
|
wenzelm@6383
|
101 |
|
wenzelm@6383
|
102 |
fun prepare_typedef prep_term no_def name (t, vs, mx) raw_set thy =
|
wenzelm@4866
|
103 |
let
|
paulson@4970
|
104 |
val _ = Theory.requires thy "Set" "typedefs";
|
wenzelm@6383
|
105 |
val sign = Theory.sign_of thy;
|
wenzelm@4866
|
106 |
val full_name = Sign.full_name sign;
|
wenzelm@4866
|
107 |
|
wenzelm@4866
|
108 |
(*rhs*)
|
wenzelm@6383
|
109 |
val cset = prep_term sign vs raw_set;
|
wenzelm@6383
|
110 |
val {T = setT, t = set, ...} = Thm.rep_cterm cset;
|
wenzelm@4866
|
111 |
val rhs_tfrees = term_tfrees set;
|
wenzelm@4866
|
112 |
val oldT = HOLogic.dest_setT setT handle TYPE _ =>
|
wenzelm@4866
|
113 |
error ("Not a set type: " ^ quote (Sign.string_of_typ sign setT));
|
wenzelm@4866
|
114 |
|
wenzelm@4866
|
115 |
(*lhs*)
|
wenzelm@6383
|
116 |
val lhs_tfrees = map (fn v => (v, if_none (assoc (rhs_tfrees, v)) HOLogic.termS)) vs;
|
wenzelm@4866
|
117 |
val tname = Syntax.type_name t mx;
|
wenzelm@4866
|
118 |
val newT = Type (full_name tname, map TFree lhs_tfrees);
|
wenzelm@4866
|
119 |
|
wenzelm@4866
|
120 |
val Rep_name = "Rep_" ^ name;
|
wenzelm@4866
|
121 |
val Abs_name = "Abs_" ^ name;
|
wenzelm@4866
|
122 |
val setC = Const (full_name name, setT);
|
wenzelm@4866
|
123 |
val RepC = Const (full_name Rep_name, newT --> oldT);
|
wenzelm@4866
|
124 |
val AbsC = Const (full_name Abs_name, oldT --> newT);
|
wenzelm@4866
|
125 |
val x_new = Free ("x", newT);
|
wenzelm@4866
|
126 |
val y_old = Free ("y", oldT);
|
berghofe@5180
|
127 |
val set' = if no_def then set else setC;
|
wenzelm@4866
|
128 |
|
wenzelm@4866
|
129 |
(*axioms*)
|
berghofe@5180
|
130 |
val rep_type = HOLogic.mk_Trueprop (HOLogic.mk_mem (RepC $ x_new, set'));
|
wenzelm@4866
|
131 |
val rep_type_inv = HOLogic.mk_Trueprop (HOLogic.mk_eq (AbsC $ (RepC $ x_new), x_new));
|
berghofe@5180
|
132 |
val abs_type_inv = Logic.mk_implies (HOLogic.mk_Trueprop (HOLogic.mk_mem (y_old, set')),
|
wenzelm@4866
|
133 |
HOLogic.mk_Trueprop (HOLogic.mk_eq (RepC $ (AbsC $ y_old), y_old)));
|
wenzelm@4866
|
134 |
|
wenzelm@6383
|
135 |
(*theory extender*)
|
wenzelm@6383
|
136 |
fun do_typedef theory =
|
wenzelm@6383
|
137 |
theory
|
wenzelm@6383
|
138 |
|> Theory.assert_super thy
|
wenzelm@6383
|
139 |
|> add_typedecls [(t, vs, mx)]
|
wenzelm@6383
|
140 |
|> Theory.add_consts_i
|
wenzelm@6383
|
141 |
((if no_def then [] else [(name, setT, NoSyn)]) @
|
wenzelm@6383
|
142 |
[(Rep_name, newT --> oldT, NoSyn),
|
wenzelm@6383
|
143 |
(Abs_name, oldT --> newT, NoSyn)])
|
wenzelm@9315
|
144 |
|> (if no_def then I else (#1 oo (PureThy.add_defs_i false o map Thm.no_attributes))
|
wenzelm@6383
|
145 |
[Logic.mk_defpair (setC, set)])
|
wenzelm@8428
|
146 |
|> (#1 oo (PureThy.add_axioms_i o map Thm.no_attributes))
|
wenzelm@6383
|
147 |
[(Rep_name, rep_type),
|
wenzelm@6383
|
148 |
(Rep_name ^ "_inverse", rep_type_inv),
|
wenzelm@6383
|
149 |
(Abs_name ^ "_inverse", abs_type_inv)]
|
wenzelm@6383
|
150 |
handle ERROR => err_in_typedef name;
|
wenzelm@6383
|
151 |
|
wenzelm@4866
|
152 |
|
wenzelm@4866
|
153 |
(* errors *)
|
wenzelm@4866
|
154 |
|
wenzelm@4866
|
155 |
fun show_names pairs = commas_quote (map fst pairs);
|
wenzelm@4866
|
156 |
|
wenzelm@4866
|
157 |
val illegal_vars =
|
wenzelm@4866
|
158 |
if null (term_vars set) andalso null (term_tvars set) then []
|
wenzelm@4866
|
159 |
else ["Illegal schematic variable(s) on rhs"];
|
wenzelm@4866
|
160 |
|
wenzelm@4866
|
161 |
val dup_lhs_tfrees =
|
wenzelm@4866
|
162 |
(case duplicates lhs_tfrees of [] => []
|
wenzelm@4866
|
163 |
| dups => ["Duplicate type variables on lhs: " ^ show_names dups]);
|
wenzelm@4866
|
164 |
|
wenzelm@4866
|
165 |
val extra_rhs_tfrees =
|
wenzelm@4866
|
166 |
(case gen_rems (op =) (rhs_tfrees, lhs_tfrees) of [] => []
|
wenzelm@4866
|
167 |
| extras => ["Extra type variables on rhs: " ^ show_names extras]);
|
wenzelm@4866
|
168 |
|
wenzelm@4866
|
169 |
val illegal_frees =
|
wenzelm@4866
|
170 |
(case term_frees set of [] => []
|
wenzelm@4866
|
171 |
| xs => ["Illegal variables on rhs: " ^ show_names (map dest_Free xs)]);
|
wenzelm@4866
|
172 |
|
wenzelm@4866
|
173 |
val errs = illegal_vars @ dup_lhs_tfrees @ extra_rhs_tfrees @ illegal_frees;
|
wenzelm@4866
|
174 |
in
|
wenzelm@6383
|
175 |
if null errs then () else error (cat_lines errs);
|
wenzelm@6383
|
176 |
(cset, do_typedef)
|
wenzelm@6383
|
177 |
end handle ERROR => err_in_typedef name;
|
wenzelm@4866
|
178 |
|
wenzelm@4866
|
179 |
|
wenzelm@6383
|
180 |
(* add_typedef interfaces *)
|
wenzelm@4866
|
181 |
|
wenzelm@6383
|
182 |
fun gen_add_typedef prep_term no_def name typ set names thms tac thy =
|
wenzelm@6383
|
183 |
let
|
wenzelm@6383
|
184 |
val (cset, do_typedef) = prepare_typedef prep_term no_def name typ set thy;
|
wenzelm@6383
|
185 |
val result = prove_nonempty thy cset (names, thms, tac);
|
wenzelm@6383
|
186 |
in check_nonempty cset result; thy |> do_typedef end;
|
wenzelm@4866
|
187 |
|
berghofe@5180
|
188 |
val add_typedef = gen_add_typedef read_term false;
|
berghofe@5180
|
189 |
val add_typedef_i = gen_add_typedef cert_term false;
|
berghofe@5180
|
190 |
val add_typedef_i_no_def = gen_add_typedef cert_term true;
|
wenzelm@4866
|
191 |
|
wenzelm@4866
|
192 |
|
wenzelm@6383
|
193 |
(* typedef_proof interface *)
|
wenzelm@6383
|
194 |
|
wenzelm@6383
|
195 |
fun typedef_attribute cset do_typedef (thy, thm) =
|
paulson@9969
|
196 |
(check_nonempty cset (thm RS (some_eq_ex RS iffD2)); (thy |> do_typedef, thm));
|
wenzelm@6357
|
197 |
|
wenzelm@6729
|
198 |
fun gen_typedef_proof prep_term ((name, typ, set), comment) int thy =
|
wenzelm@6383
|
199 |
let
|
wenzelm@6383
|
200 |
val (cset, do_typedef) = prepare_typedef prep_term false name typ set thy;
|
wenzelm@7481
|
201 |
val goal = Thm.term_of (goal_nonempty true cset);
|
wenzelm@6383
|
202 |
in
|
wenzelm@6383
|
203 |
thy
|
wenzelm@6935
|
204 |
|> IsarThy.theorem_i (("", [typedef_attribute cset do_typedef], (goal, ([], []))), comment) int
|
wenzelm@6383
|
205 |
end;
|
wenzelm@6383
|
206 |
|
wenzelm@6383
|
207 |
val typedef_proof = gen_typedef_proof read_term;
|
wenzelm@6383
|
208 |
val typedef_proof_i = gen_typedef_proof cert_term;
|
wenzelm@6357
|
209 |
|
wenzelm@6357
|
210 |
|
wenzelm@6383
|
211 |
|
wenzelm@6383
|
212 |
(** outer syntax **)
|
wenzelm@6383
|
213 |
|
wenzelm@6723
|
214 |
local structure P = OuterParse and K = OuterSyntax.Keyword in
|
wenzelm@6383
|
215 |
|
wenzelm@6357
|
216 |
val typedeclP =
|
wenzelm@6723
|
217 |
OuterSyntax.command "typedecl" "HOL type declaration" K.thy_decl
|
wenzelm@7152
|
218 |
(P.type_args -- P.name -- P.opt_infix --| P.marg_comment >> (fn ((vs, t), mx) =>
|
wenzelm@6357
|
219 |
Toplevel.theory (add_typedecls [(t, vs, mx)])));
|
wenzelm@6357
|
220 |
|
wenzelm@6723
|
221 |
|
wenzelm@6383
|
222 |
val typedef_proof_decl =
|
wenzelm@6723
|
223 |
Scan.option (P.$$$ "(" |-- P.name --| P.$$$ ")") --
|
wenzelm@6729
|
224 |
(P.type_args -- P.name) -- P.opt_infix -- (P.$$$ "=" |-- P.term) -- P.marg_comment;
|
wenzelm@6357
|
225 |
|
wenzelm@6729
|
226 |
fun mk_typedef_proof ((((opt_name, (vs, t)), mx), A), comment) =
|
wenzelm@6729
|
227 |
typedef_proof ((if_none opt_name (Syntax.type_name t mx), (t, vs, mx), A), comment);
|
wenzelm@6357
|
228 |
|
wenzelm@6357
|
229 |
val typedefP =
|
wenzelm@6723
|
230 |
OuterSyntax.command "typedef" "HOL type definition (requires non-emptiness proof)" K.thy_goal
|
wenzelm@6383
|
231 |
(typedef_proof_decl >> (Toplevel.print oo (Toplevel.theory_to_proof o mk_typedef_proof)));
|
wenzelm@6357
|
232 |
|
wenzelm@6723
|
233 |
|
wenzelm@6357
|
234 |
val _ = OuterSyntax.add_parsers [typedeclP, typedefP];
|
wenzelm@6357
|
235 |
|
wenzelm@4866
|
236 |
end;
|
wenzelm@6383
|
237 |
|
wenzelm@6383
|
238 |
|
wenzelm@6383
|
239 |
end;
|