| author | wenzelm | 
| Tue, 25 Oct 2005 18:18:58 +0200 | |
| changeset 17987 | f8b45ab11400 | 
| parent 17895 | 6274b426594b | 
| child 18358 | 0a733e11021a | 
| 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  | 
11  | 
val add_specification_i: string option -> (bstring * xstring * bool) list ->  | 
|
12  | 
theory attribute  | 
|
| 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)  | 
|
35  | 
val cname_full = Sign.intern_const (sign_of thy) cname  | 
|
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)  | 
|
41  | 
val (thy',thms) = 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  | 
|
52  | 
val (thy',thms) = PureThy.add_axioms_i [((axname,HOLogic.mk_Trueprop tm),[])] thy  | 
|
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)  | 
|
61  | 
val cname_full = Sign.intern_const (sign_of thy) cname  | 
|
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  | 
|
| 
 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 
skalberg 
parents: 
14167 
diff
changeset
 | 
85  | 
fun add_specification_i axiomatic cos arg =  | 
| 
14121
 
d2a0fd183f5f
Added handling of free variables (provided they are of sort HOL.type).
 
skalberg 
parents: 
14118 
diff
changeset
 | 
86  | 
arg |> apsnd freezeT  | 
| 17336 | 87  | 
|> proc_exprop axiomatic cos  | 
88  | 
|> apsnd standard  | 
|
| 14115 | 89  | 
|
90  | 
(* Collect all intances of constants in term *)  | 
|
91  | 
||
92  | 
fun collect_consts ( t $ u,tms) = collect_consts (u,collect_consts (t,tms))  | 
|
93  | 
| collect_consts ( Abs(_,_,t),tms) = collect_consts (t,tms)  | 
|
94  | 
| collect_consts (tm as Const _,tms) = gen_ins (op aconv) (tm,tms)  | 
|
95  | 
| collect_consts ( _,tms) = tms  | 
|
96  | 
||
97  | 
(* Complementing Type.varify... *)  | 
|
98  | 
||
99  | 
fun unvarify t fmap =  | 
|
100  | 
let  | 
|
| 17336 | 101  | 
val fmap' = map Library.swap fmap  | 
| 17377 | 102  | 
fun unthaw (f as (a, S)) =  | 
103  | 
(case AList.lookup (op =) fmap' a of  | 
|
| 17336 | 104  | 
NONE => TVar f  | 
| 17377 | 105  | 
| SOME (b, _) => TFree (b, S))  | 
| 14115 | 106  | 
in  | 
| 17336 | 107  | 
map_term_types (map_type_tvar unthaw) t  | 
| 14115 | 108  | 
end  | 
109  | 
||
110  | 
(* The syntactic meddling needed to setup add_specification for work *)  | 
|
111  | 
||
| 17336 | 112  | 
fun process_spec axiomatic cos alt_props thy =  | 
| 14115 | 113  | 
let  | 
| 17336 | 114  | 
fun zip3 [] [] [] = []  | 
115  | 
| zip3 (x::xs) (y::ys) (z::zs) = (x,y,z)::zip3 xs ys zs  | 
|
116  | 
| zip3 _ _ _ = error "SpecificationPackage.process_spec internal error"  | 
|
| 14167 | 117  | 
|
| 17336 | 118  | 
fun myfoldr f [x] = x  | 
119  | 
| myfoldr f (x::xs) = f (x,myfoldr f xs)  | 
|
120  | 
| myfoldr f [] = error "SpecificationPackage.process_spec internal error"  | 
|
| 
14164
 
8c3fab596219
Allowed for splitting the specification over several lemmas.
 
skalberg 
parents: 
14122 
diff
changeset
 | 
121  | 
|
| 17895 | 122  | 
fun typ_equiv t u = Sign.typ_instance thy (t,u) andalso Sign.typ_instance thy (u,t);  | 
| 
14164
 
8c3fab596219
Allowed for splitting the specification over several lemmas.
 
skalberg 
parents: 
14122 
diff
changeset
 | 
123  | 
|
| 17895 | 124  | 
val rew_imps = alt_props |>  | 
125  | 
map (ObjectLogic.atomize_cterm thy o Thm.read_cterm thy o rpair Term.propT o snd)  | 
|
126  | 
val props' = rew_imps |>  | 
|
127  | 
map (HOLogic.dest_Trueprop o term_of o snd o dest_equals o cprop_of)  | 
|
| 14166 | 128  | 
|
| 17336 | 129  | 
fun proc_single prop =  | 
130  | 
let  | 
|
131  | 
val frees = Term.term_frees prop  | 
|
132  | 
val tsig = Sign.tsig_of (sign_of thy)  | 
|
133  | 
val _ = assert (forall (fn v => Type.of_sort tsig (type_of v,HOLogic.typeS)) frees)  | 
|
134  | 
"Specificaton: Only free variables of sort 'type' allowed"  | 
|
135  | 
val prop_closed = foldr (fn ((vname,T),prop) => HOLogic.mk_all (vname,T,prop)) prop (map dest_Free frees)  | 
|
136  | 
in  | 
|
137  | 
(prop_closed,frees)  | 
|
138  | 
end  | 
|
| 14166 | 139  | 
|
| 17336 | 140  | 
val props'' = map proc_single props'  | 
141  | 
val frees = map snd props''  | 
|
142  | 
val prop = myfoldr HOLogic.mk_conj (map fst props'')  | 
|
| 17895 | 143  | 
val cprop = cterm_of thy (HOLogic.mk_Trueprop prop)  | 
| 
14164
 
8c3fab596219
Allowed for splitting the specification over several lemmas.
 
skalberg 
parents: 
14122 
diff
changeset
 | 
144  | 
|
| 17336 | 145  | 
val (prop_thawed,vmap) = Type.varify (prop,[])  | 
146  | 
val thawed_prop_consts = collect_consts (prop_thawed,[])  | 
|
147  | 
val (altcos,overloaded) = Library.split_list cos  | 
|
148  | 
val (names,sconsts) = Library.split_list altcos  | 
|
| 17895 | 149  | 
val consts = map (term_of o Thm.read_cterm thy o rpair TypeInfer.logicT) sconsts  | 
| 17336 | 150  | 
val _ = assert (not (Library.exists (not o Term.is_Const) consts))  | 
151  | 
"Specification: Non-constant found as parameter"  | 
|
| 14166 | 152  | 
|
| 17336 | 153  | 
fun proc_const c =  | 
154  | 
let  | 
|
155  | 
val c' = fst (Type.varify (c,[]))  | 
|
156  | 
val (cname,ctyp) = dest_Const c'  | 
|
157  | 
in  | 
|
158  | 
case List.filter (fn t => let val (name,typ) = dest_Const t  | 
|
159  | 
in name = cname andalso typ_equiv typ ctyp  | 
|
160  | 
end) thawed_prop_consts of  | 
|
| 17895 | 161  | 
                    [] => error ("Specification: No suitable instances of constant \"" ^ (Sign.string_of_term thy c) ^ "\" found")
 | 
| 17336 | 162  | 
| [cf] => unvarify cf vmap  | 
| 17895 | 163  | 
                  | _ => error ("Specification: Several variations of \"" ^ (Sign.string_of_term thy c) ^ "\" found (try applying explicit type constraints)")
 | 
| 17336 | 164  | 
end  | 
165  | 
val proc_consts = map proc_const consts  | 
|
166  | 
fun mk_exist (c,prop) =  | 
|
167  | 
let  | 
|
168  | 
val T = type_of c  | 
|
169  | 
val cname = Sign.base_name (fst (dest_Const c))  | 
|
170  | 
val vname = if Syntax.is_identifier cname  | 
|
171  | 
then cname  | 
|
172  | 
else "x"  | 
|
173  | 
in  | 
|
174  | 
HOLogic.exists_const T $ Abs(vname,T,Term.abstract_over (c,prop))  | 
|
175  | 
end  | 
|
176  | 
val ex_prop = foldr mk_exist prop proc_consts  | 
|
177  | 
val cnames = map (fst o dest_Const) proc_consts  | 
|
178  | 
fun post_process (arg as (thy,thm)) =  | 
|
179  | 
let  | 
|
| 17895 | 180  | 
fun inst_all thy (thm,v) =  | 
| 17336 | 181  | 
let  | 
| 17895 | 182  | 
val cv = cterm_of thy v  | 
| 17336 | 183  | 
val cT = ctyp_of_term cv  | 
184  | 
val spec' = instantiate' [SOME cT] [NONE,SOME cv] spec  | 
|
185  | 
in  | 
|
186  | 
thm RS spec'  | 
|
187  | 
end  | 
|
188  | 
fun remove_alls frees thm = Library.foldl (inst_all (sign_of_thm thm)) (thm,frees)  | 
|
189  | 
fun process_single ((name,atts),rew_imp,frees) args =  | 
|
190  | 
let  | 
|
191  | 
fun undo_imps thm =  | 
|
192  | 
equal_elim (symmetric rew_imp) thm  | 
|
| 
14164
 
8c3fab596219
Allowed for splitting the specification over several lemmas.
 
skalberg 
parents: 
14122 
diff
changeset
 | 
193  | 
|
| 17336 | 194  | 
fun apply_atts arg = Thm.apply_attributes (arg,map (Attrib.global_attribute thy) atts)  | 
195  | 
fun add_final (arg as (thy,thm)) =  | 
|
196  | 
if name = ""  | 
|
197  | 
then arg  | 
|
198  | 
                            else (writeln ("  " ^ name ^ ": " ^ (string_of_thm thm));
 | 
|
199  | 
PureThy.store_thm((name,thm),[]) thy)  | 
|
200  | 
in  | 
|
201  | 
args |> apsnd (remove_alls frees)  | 
|
202  | 
|> apsnd undo_imps  | 
|
203  | 
|> apsnd standard  | 
|
204  | 
|> apply_atts  | 
|
205  | 
|> add_final  | 
|
206  | 
end  | 
|
| 
14164
 
8c3fab596219
Allowed for splitting the specification over several lemmas.
 
skalberg 
parents: 
14122 
diff
changeset
 | 
207  | 
|
| 17336 | 208  | 
fun process_all [proc_arg] args =  | 
209  | 
process_single proc_arg args  | 
|
210  | 
| process_all (proc_arg::rest) (thy,thm) =  | 
|
211  | 
let  | 
|
212  | 
val single_th = thm RS conjunct1  | 
|
213  | 
val rest_th = thm RS conjunct2  | 
|
214  | 
val (thy',_) = process_single proc_arg (thy,single_th)  | 
|
215  | 
in  | 
|
216  | 
process_all rest (thy',rest_th)  | 
|
217  | 
end  | 
|
218  | 
| process_all [] _ = error "SpecificationPackage.process_spec internal error"  | 
|
219  | 
val alt_names = map fst alt_props  | 
|
220  | 
val _ = if exists (fn(name,_) => not (name = "")) alt_names  | 
|
221  | 
then writeln "specification"  | 
|
222  | 
else ()  | 
|
223  | 
in  | 
|
224  | 
arg |> apsnd freezeT  | 
|
225  | 
|> process_all (zip3 alt_names rew_imps frees)  | 
|
226  | 
end  | 
|
| 14115 | 227  | 
in  | 
| 17336 | 228  | 
IsarThy.theorem_i Drule.internalK  | 
229  | 
        ("", [add_specification_i axiomatic (zip3 names cnames overloaded), post_process])
 | 
|
230  | 
(HOLogic.mk_Trueprop ex_prop, ([], [])) thy  | 
|
| 14115 | 231  | 
end  | 
232  | 
||
| 17336 | 233  | 
|
| 14115 | 234  | 
(* outer syntax *)  | 
235  | 
||
| 17057 | 236  | 
local structure P = OuterParse and K = OuterKeyword in  | 
| 14115 | 237  | 
|
| 
14116
 
63337d8e6e0f
Added optional theorem names for the constant definitions added during
 
skalberg 
parents: 
14115 
diff
changeset
 | 
238  | 
val opt_name = Scan.optional (P.name --| P.$$$ ":") ""  | 
| 17336 | 239  | 
val opt_overloaded = P.opt_keyword "overloaded";  | 
| 
14116
 
63337d8e6e0f
Added optional theorem names for the constant definitions added during
 
skalberg 
parents: 
14115 
diff
changeset
 | 
240  | 
|
| 14115 | 241  | 
val specification_decl =  | 
| 17336 | 242  | 
  P.$$$ "(" |-- Scan.repeat1 (opt_name -- P.term -- opt_overloaded) --| P.$$$ ")" --
 | 
243  | 
Scan.repeat1 (P.opt_thm_name ":" -- P.prop)  | 
|
| 14115 | 244  | 
|
245  | 
val specificationP =  | 
|
246  | 
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
 | 
247  | 
(specification_decl >> (fn (cos,alt_props) =>  | 
| 17336 | 248  | 
Toplevel.print o (Toplevel.theory_to_proof  | 
249  | 
(process_spec NONE cos alt_props))))  | 
|
| 
14222
 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 
skalberg 
parents: 
14167 
diff
changeset
 | 
250  | 
|
| 
 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 
skalberg 
parents: 
14167 
diff
changeset
 | 
251  | 
val ax_specification_decl =  | 
| 
 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 
skalberg 
parents: 
14167 
diff
changeset
 | 
252  | 
P.name --  | 
| 
 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 
skalberg 
parents: 
14167 
diff
changeset
 | 
253  | 
    (P.$$$ "(" |-- Scan.repeat1 (opt_name -- P.term -- opt_overloaded) --| P.$$$ ")" --
 | 
| 17336 | 254  | 
Scan.repeat1 (P.opt_thm_name ":" -- P.prop))  | 
| 14115 | 255  | 
|
| 
14222
 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 
skalberg 
parents: 
14167 
diff
changeset
 | 
256  | 
val ax_specificationP =  | 
| 
 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 
skalberg 
parents: 
14167 
diff
changeset
 | 
257  | 
OuterSyntax.command "ax_specification" "define constants by specification" K.thy_goal  | 
| 
 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 
skalberg 
parents: 
14167 
diff
changeset
 | 
258  | 
(ax_specification_decl >> (fn (axname,(cos,alt_props)) =>  | 
| 17336 | 259  | 
Toplevel.print o (Toplevel.theory_to_proof  | 
260  | 
(process_spec (SOME axname) cos alt_props))))  | 
|
| 
14222
 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 
skalberg 
parents: 
14167 
diff
changeset
 | 
261  | 
|
| 
 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 
skalberg 
parents: 
14167 
diff
changeset
 | 
262  | 
val _ = OuterSyntax.add_parsers [specificationP,ax_specificationP]  | 
| 14115 | 263  | 
|
264  | 
end  | 
|
265  | 
||
266  | 
||
267  | 
end  |