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