author | wenzelm |
Thu, 21 Jun 2007 17:28:53 +0200 | |
changeset 23463 | 9953ff53cc64 |
parent 22902 | ac833b4bb7ee |
child 23591 | d32a85385e17 |
permissions | -rw-r--r-- |
14115 | 1 |
(* Title: HOL/Tools/specification_package.ML |
2 |
ID: $Id$ |
|
3 |
Author: Sebastian Skalberg, TU Muenchen |
|
4 |
||
5 |
Package for defining constants by specification. |
|
6 |
*) |
|
7 |
||
8 |
signature SPECIFICATION_PACKAGE = |
|
9 |
sig |
|
17336 | 10 |
val quiet_mode: bool ref |
18729 | 11 |
val add_specification: string option -> (bstring * xstring * bool) list -> |
12 |
theory * thm -> theory * thm |
|
14115 | 13 |
end |
14 |
||
17895 | 15 |
structure SpecificationPackage: SPECIFICATION_PACKAGE = |
14115 | 16 |
struct |
17 |
||
18 |
(* messages *) |
|
19 |
||
20 |
val quiet_mode = ref false |
|
21 |
fun message s = if ! quiet_mode then () else writeln s |
|
22 |
||
23 |
||
24 |
(* Actual code *) |
|
25 |
||
14222
1e61f23fd359
Added axiomatic specifications (ax_specification).
skalberg
parents:
14167
diff
changeset
|
26 |
local |
17895 | 27 |
val exE_some = thm "exE_some"; |
28 |
||
14222
1e61f23fd359
Added axiomatic specifications (ax_specification).
skalberg
parents:
14167
diff
changeset
|
29 |
fun mk_definitional [] arg = arg |
1e61f23fd359
Added axiomatic specifications (ax_specification).
skalberg
parents:
14167
diff
changeset
|
30 |
| mk_definitional ((thname,cname,covld)::cos) (thy,thm) = |
17336 | 31 |
case HOLogic.dest_Trueprop (concl_of thm) of |
32 |
Const("Ex",_) $ P => |
|
33 |
let |
|
34 |
val ctype = domain_type (type_of P) |
|
22578 | 35 |
val cname_full = Sign.intern_const thy cname |
17336 | 36 |
val cdefname = if thname = "" |
37 |
then Thm.def_name (Sign.base_name cname) |
|
38 |
else thname |
|
39 |
val def_eq = Logic.mk_equals (Const(cname_full,ctype), |
|
40 |
HOLogic.choice_const ctype $ P) |
|
18358 | 41 |
val (thms, thy') = PureThy.add_defs_i covld [((cdefname,def_eq),[])] thy |
17895 | 42 |
val thm' = [thm,hd thms] MRS exE_some |
17336 | 43 |
in |
44 |
mk_definitional cos (thy',thm') |
|
45 |
end |
|
46 |
| _ => raise THM ("Internal error: Bad specification theorem",0,[thm]) |
|
14222
1e61f23fd359
Added axiomatic specifications (ax_specification).
skalberg
parents:
14167
diff
changeset
|
47 |
|
1e61f23fd359
Added axiomatic specifications (ax_specification).
skalberg
parents:
14167
diff
changeset
|
48 |
fun mk_axiomatic axname cos arg = |
17336 | 49 |
let |
50 |
fun process [] (thy,tm) = |
|
51 |
let |
|
18377 | 52 |
val (thms, thy') = PureThy.add_axioms_i [((axname,HOLogic.mk_Trueprop tm),[])] thy |
17336 | 53 |
in |
54 |
(thy',hd thms) |
|
55 |
end |
|
56 |
| process ((thname,cname,covld)::cos) (thy,tm) = |
|
57 |
case tm of |
|
58 |
Const("Ex",_) $ P => |
|
59 |
let |
|
60 |
val ctype = domain_type (type_of P) |
|
22578 | 61 |
val cname_full = Sign.intern_const thy cname |
17336 | 62 |
val cdefname = if thname = "" |
63 |
then Thm.def_name (Sign.base_name cname) |
|
64 |
else thname |
|
65 |
val co = Const(cname_full,ctype) |
|
66 |
val thy' = Theory.add_finals_i covld [co] thy |
|
67 |
val tm' = case P of |
|
68 |
Abs(_, _, bodt) => subst_bound (co, bodt) |
|
69 |
| _ => P $ co |
|
70 |
in |
|
71 |
process cos (thy',tm') |
|
72 |
end |
|
73 |
| _ => raise TERM ("Internal error: Bad specification theorem",[tm]) |
|
74 |
in |
|
75 |
process cos arg |
|
76 |
end |
|
14115 | 77 |
|
14222
1e61f23fd359
Added axiomatic specifications (ax_specification).
skalberg
parents:
14167
diff
changeset
|
78 |
in |
1e61f23fd359
Added axiomatic specifications (ax_specification).
skalberg
parents:
14167
diff
changeset
|
79 |
fun proc_exprop axiomatic cos arg = |
1e61f23fd359
Added axiomatic specifications (ax_specification).
skalberg
parents:
14167
diff
changeset
|
80 |
case axiomatic of |
17336 | 81 |
SOME axname => mk_axiomatic axname cos (apsnd (HOLogic.dest_Trueprop o concl_of) arg) |
15531 | 82 |
| NONE => mk_definitional cos arg |
14222
1e61f23fd359
Added axiomatic specifications (ax_specification).
skalberg
parents:
14167
diff
changeset
|
83 |
end |
1e61f23fd359
Added axiomatic specifications (ax_specification).
skalberg
parents:
14167
diff
changeset
|
84 |
|
18729 | 85 |
fun add_specification axiomatic cos arg = |
19876 | 86 |
arg |> apsnd Thm.freezeT |
17336 | 87 |
|> proc_exprop axiomatic cos |
88 |
|> apsnd standard |
|
14115 | 89 |
|
18729 | 90 |
|
14115 | 91 |
(* Collect all intances of constants in term *) |
92 |
||
93 |
fun collect_consts ( t $ u,tms) = collect_consts (u,collect_consts (t,tms)) |
|
94 |
| collect_consts ( Abs(_,_,t),tms) = collect_consts (t,tms) |
|
18921 | 95 |
| collect_consts (tm as Const _,tms) = insert (op aconv) tm tms |
14115 | 96 |
| collect_consts ( _,tms) = tms |
97 |
||
98 |
(* Complementing Type.varify... *) |
|
99 |
||
100 |
fun unvarify t fmap = |
|
101 |
let |
|
17336 | 102 |
val fmap' = map Library.swap fmap |
17377 | 103 |
fun unthaw (f as (a, S)) = |
104 |
(case AList.lookup (op =) fmap' a of |
|
17336 | 105 |
NONE => TVar f |
17377 | 106 |
| SOME (b, _) => TFree (b, S)) |
14115 | 107 |
in |
20548
8ef25fe585a8
renamed Term.map_term_types to Term.map_types (cf. Term.fold_types);
wenzelm
parents:
20344
diff
changeset
|
108 |
map_types (map_type_tvar unthaw) t |
14115 | 109 |
end |
110 |
||
111 |
(* The syntactic meddling needed to setup add_specification for work *) |
|
112 |
||
17336 | 113 |
fun process_spec axiomatic cos alt_props thy = |
14115 | 114 |
let |
17336 | 115 |
fun zip3 [] [] [] = [] |
116 |
| zip3 (x::xs) (y::ys) (z::zs) = (x,y,z)::zip3 xs ys zs |
|
117 |
| zip3 _ _ _ = error "SpecificationPackage.process_spec internal error" |
|
14167 | 118 |
|
17336 | 119 |
fun myfoldr f [x] = x |
120 |
| myfoldr f (x::xs) = f (x,myfoldr f xs) |
|
121 |
| myfoldr f [] = error "SpecificationPackage.process_spec internal error" |
|
14164
8c3fab596219
Allowed for splitting the specification over several lemmas.
skalberg
parents:
14122
diff
changeset
|
122 |
|
17895 | 123 |
val rew_imps = alt_props |> |
18776 | 124 |
map (ObjectLogic.atomize_cterm o Thm.read_cterm thy o rpair Term.propT o snd) |
17895 | 125 |
val props' = rew_imps |> |
22902
ac833b4bb7ee
moved some Drule operations to Thm (see more_thm.ML);
wenzelm
parents:
22709
diff
changeset
|
126 |
map (HOLogic.dest_Trueprop o term_of o snd o Thm.dest_equals o cprop_of) |
14166 | 127 |
|
17336 | 128 |
fun proc_single prop = |
129 |
let |
|
130 |
val frees = Term.term_frees prop |
|
21858
05f57309170c
avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
wenzelm
parents:
21434
diff
changeset
|
131 |
val _ = forall (fn v => Sign.of_sort thy (type_of v,HOLogic.typeS)) frees |
05f57309170c
avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
wenzelm
parents:
21434
diff
changeset
|
132 |
orelse error "Specificaton: Only free variables of sort 'type' allowed" |
17336 | 133 |
val prop_closed = foldr (fn ((vname,T),prop) => HOLogic.mk_all (vname,T,prop)) prop (map dest_Free frees) |
134 |
in |
|
135 |
(prop_closed,frees) |
|
136 |
end |
|
14166 | 137 |
|
17336 | 138 |
val props'' = map proc_single props' |
139 |
val frees = map snd props'' |
|
140 |
val prop = myfoldr HOLogic.mk_conj (map fst props'') |
|
17895 | 141 |
val cprop = cterm_of thy (HOLogic.mk_Trueprop prop) |
14164
8c3fab596219
Allowed for splitting the specification over several lemmas.
skalberg
parents:
14122
diff
changeset
|
142 |
|
21116 | 143 |
val (vmap, prop_thawed) = Type.varify [] prop |
17336 | 144 |
val thawed_prop_consts = collect_consts (prop_thawed,[]) |
145 |
val (altcos,overloaded) = Library.split_list cos |
|
146 |
val (names,sconsts) = Library.split_list altcos |
|
22709 | 147 |
val consts = map (term_of o Thm.read_cterm thy o rpair dummyT) sconsts |
21858
05f57309170c
avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
wenzelm
parents:
21434
diff
changeset
|
148 |
val _ = not (Library.exists (not o Term.is_Const) consts) |
05f57309170c
avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
wenzelm
parents:
21434
diff
changeset
|
149 |
orelse error "Specification: Non-constant found as parameter" |
14166 | 150 |
|
17336 | 151 |
fun proc_const c = |
152 |
let |
|
21116 | 153 |
val (_, c') = Type.varify [] c |
17336 | 154 |
val (cname,ctyp) = dest_Const c' |
155 |
in |
|
156 |
case List.filter (fn t => let val (name,typ) = dest_Const t |
|
19414 | 157 |
in name = cname andalso Sign.typ_equiv thy (typ, ctyp) |
17336 | 158 |
end) thawed_prop_consts of |
17895 | 159 |
[] => error ("Specification: No suitable instances of constant \"" ^ (Sign.string_of_term thy c) ^ "\" found") |
17336 | 160 |
| [cf] => unvarify cf vmap |
17895 | 161 |
| _ => error ("Specification: Several variations of \"" ^ (Sign.string_of_term thy c) ^ "\" found (try applying explicit type constraints)") |
17336 | 162 |
end |
163 |
val proc_consts = map proc_const consts |
|
164 |
fun mk_exist (c,prop) = |
|
165 |
let |
|
166 |
val T = type_of c |
|
167 |
val cname = Sign.base_name (fst (dest_Const c)) |
|
168 |
val vname = if Syntax.is_identifier cname |
|
169 |
then cname |
|
170 |
else "x" |
|
171 |
in |
|
172 |
HOLogic.exists_const T $ Abs(vname,T,Term.abstract_over (c,prop)) |
|
173 |
end |
|
174 |
val ex_prop = foldr mk_exist prop proc_consts |
|
175 |
val cnames = map (fst o dest_Const) proc_consts |
|
176 |
fun post_process (arg as (thy,thm)) = |
|
177 |
let |
|
17895 | 178 |
fun inst_all thy (thm,v) = |
17336 | 179 |
let |
17895 | 180 |
val cv = cterm_of thy v |
17336 | 181 |
val cT = ctyp_of_term cv |
182 |
val spec' = instantiate' [SOME cT] [NONE,SOME cv] spec |
|
183 |
in |
|
184 |
thm RS spec' |
|
185 |
end |
|
22578 | 186 |
fun remove_alls frees thm = |
187 |
Library.foldl (inst_all (Thm.theory_of_thm thm)) (thm,frees) |
|
17336 | 188 |
fun process_single ((name,atts),rew_imp,frees) args = |
189 |
let |
|
190 |
fun undo_imps thm = |
|
191 |
equal_elim (symmetric rew_imp) thm |
|
14164
8c3fab596219
Allowed for splitting the specification over several lemmas.
skalberg
parents:
14122
diff
changeset
|
192 |
|
18358 | 193 |
fun add_final (arg as (thy, thm)) = |
17336 | 194 |
if name = "" |
18358 | 195 |
then arg |> Library.swap |
17336 | 196 |
else (writeln (" " ^ name ^ ": " ^ (string_of_thm thm)); |
18358 | 197 |
PureThy.store_thm ((name, thm), []) thy) |
17336 | 198 |
in |
199 |
args |> apsnd (remove_alls frees) |
|
200 |
|> apsnd undo_imps |
|
201 |
|> apsnd standard |
|
18729 | 202 |
|> Thm.theory_attributes (map (Attrib.attribute thy) atts) |
17336 | 203 |
|> add_final |
18358 | 204 |
|> Library.swap |
17336 | 205 |
end |
14164
8c3fab596219
Allowed for splitting the specification over several lemmas.
skalberg
parents:
14122
diff
changeset
|
206 |
|
17336 | 207 |
fun process_all [proc_arg] args = |
208 |
process_single proc_arg args |
|
209 |
| process_all (proc_arg::rest) (thy,thm) = |
|
210 |
let |
|
211 |
val single_th = thm RS conjunct1 |
|
212 |
val rest_th = thm RS conjunct2 |
|
213 |
val (thy',_) = process_single proc_arg (thy,single_th) |
|
214 |
in |
|
215 |
process_all rest (thy',rest_th) |
|
216 |
end |
|
217 |
| process_all [] _ = error "SpecificationPackage.process_spec internal error" |
|
218 |
val alt_names = map fst alt_props |
|
219 |
val _ = if exists (fn(name,_) => not (name = "")) alt_names |
|
220 |
then writeln "specification" |
|
221 |
else () |
|
222 |
in |
|
19876 | 223 |
arg |> apsnd Thm.freezeT |
17336 | 224 |
|> process_all (zip3 alt_names rew_imps frees) |
225 |
end |
|
20903 | 226 |
|
227 |
fun after_qed [[thm]] = ProofContext.theory (fn thy => |
|
228 |
#1 (post_process (add_specification axiomatic (zip3 names cnames overloaded) (thy, thm)))); |
|
14115 | 229 |
in |
20903 | 230 |
thy |
231 |
|> ProofContext.init |
|
21434 | 232 |
|> Proof.theorem_i NONE after_qed [[(HOLogic.mk_Trueprop ex_prop, [])]] |
20903 | 233 |
end; |
14115 | 234 |
|
17336 | 235 |
|
14115 | 236 |
(* outer syntax *) |
237 |
||
17057 | 238 |
local structure P = OuterParse and K = OuterKeyword in |
14115 | 239 |
|
14116
63337d8e6e0f
Added optional theorem names for the constant definitions added during
skalberg
parents:
14115
diff
changeset
|
240 |
val opt_name = Scan.optional (P.name --| P.$$$ ":") "" |
17336 | 241 |
val opt_overloaded = P.opt_keyword "overloaded"; |
14116
63337d8e6e0f
Added optional theorem names for the constant definitions added during
skalberg
parents:
14115
diff
changeset
|
242 |
|
14115 | 243 |
val specification_decl = |
17336 | 244 |
P.$$$ "(" |-- Scan.repeat1 (opt_name -- P.term -- opt_overloaded) --| P.$$$ ")" -- |
22101 | 245 |
Scan.repeat1 (SpecParse.opt_thm_name ":" -- P.prop) |
14115 | 246 |
|
247 |
val specificationP = |
|
248 |
OuterSyntax.command "specification" "define constants by specification" K.thy_goal |
|
14164
8c3fab596219
Allowed for splitting the specification over several lemmas.
skalberg
parents:
14122
diff
changeset
|
249 |
(specification_decl >> (fn (cos,alt_props) => |
17336 | 250 |
Toplevel.print o (Toplevel.theory_to_proof |
251 |
(process_spec NONE cos alt_props)))) |
|
14222
1e61f23fd359
Added axiomatic specifications (ax_specification).
skalberg
parents:
14167
diff
changeset
|
252 |
|
1e61f23fd359
Added axiomatic specifications (ax_specification).
skalberg
parents:
14167
diff
changeset
|
253 |
val ax_specification_decl = |
1e61f23fd359
Added axiomatic specifications (ax_specification).
skalberg
parents:
14167
diff
changeset
|
254 |
P.name -- |
1e61f23fd359
Added axiomatic specifications (ax_specification).
skalberg
parents:
14167
diff
changeset
|
255 |
(P.$$$ "(" |-- Scan.repeat1 (opt_name -- P.term -- opt_overloaded) --| P.$$$ ")" -- |
22101 | 256 |
Scan.repeat1 (SpecParse.opt_thm_name ":" -- P.prop)) |
14115 | 257 |
|
14222
1e61f23fd359
Added axiomatic specifications (ax_specification).
skalberg
parents:
14167
diff
changeset
|
258 |
val ax_specificationP = |
1e61f23fd359
Added axiomatic specifications (ax_specification).
skalberg
parents:
14167
diff
changeset
|
259 |
OuterSyntax.command "ax_specification" "define constants by specification" K.thy_goal |
1e61f23fd359
Added axiomatic specifications (ax_specification).
skalberg
parents:
14167
diff
changeset
|
260 |
(ax_specification_decl >> (fn (axname,(cos,alt_props)) => |
17336 | 261 |
Toplevel.print o (Toplevel.theory_to_proof |
262 |
(process_spec (SOME axname) cos alt_props)))) |
|
14222
1e61f23fd359
Added axiomatic specifications (ax_specification).
skalberg
parents:
14167
diff
changeset
|
263 |
|
1e61f23fd359
Added axiomatic specifications (ax_specification).
skalberg
parents:
14167
diff
changeset
|
264 |
val _ = OuterSyntax.add_parsers [specificationP,ax_specificationP] |
14115 | 265 |
|
266 |
end |
|
267 |
||
268 |
||
269 |
end |