| author | nipkow | 
| Sat, 13 Nov 2004 07:47:34 +0100 | |
| changeset 15281 | bd4611956c7b | 
| parent 14769 | b698d0b243dc | 
| child 15531 | 08c8dad8e399 | 
| 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 | |
| 10 | val quiet_mode: bool ref | |
| 14222 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 skalberg parents: 
14167diff
changeset | 11 | val add_specification_i: string option -> (bstring * xstring * bool) list | 
| 14115 | 12 | -> theory * thm -> theory * thm | 
| 13 | end | |
| 14 | ||
| 15 | structure SpecificationPackage : SPECIFICATION_PACKAGE = | |
| 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 | local | |
| 24 | val _ = Goal "[| Ex P ; c == Eps P |] ==> P c" | |
| 25 | val _ = by (Asm_simp_tac 1) | |
| 26 | val _ = by (rtac (thm "someI_ex") 1) | |
| 27 | val _ = ba 1 | |
| 28 | in | |
| 29 | val helper = Goals.result() | |
| 30 | end | |
| 31 | ||
| 32 | (* Akin to HOLogic.exists_const *) | |
| 33 | fun choice_const T = Const("Hilbert_Choice.Eps",(T-->HOLogic.boolT)-->T)
 | |
| 34 | ||
| 35 | (* Actual code *) | |
| 36 | ||
| 14222 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 skalberg parents: 
14167diff
changeset | 37 | local | 
| 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 skalberg parents: 
14167diff
changeset | 38 | fun mk_definitional [] arg = arg | 
| 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 skalberg parents: 
14167diff
changeset | 39 | | mk_definitional ((thname,cname,covld)::cos) (thy,thm) = | 
| 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 skalberg parents: 
14167diff
changeset | 40 | case HOLogic.dest_Trueprop (concl_of thm) of | 
| 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 skalberg parents: 
14167diff
changeset | 41 | 	    Const("Ex",_) $ P =>
 | 
| 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 skalberg parents: 
14167diff
changeset | 42 | let | 
| 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 skalberg parents: 
14167diff
changeset | 43 | val ctype = domain_type (type_of P) | 
| 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 skalberg parents: 
14167diff
changeset | 44 | val cname_full = Sign.intern_const (sign_of thy) cname | 
| 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 skalberg parents: 
14167diff
changeset | 45 | val cdefname = if thname = "" | 
| 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 skalberg parents: 
14167diff
changeset | 46 | then Thm.def_name (Sign.base_name cname) | 
| 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 skalberg parents: 
14167diff
changeset | 47 | else thname | 
| 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 skalberg parents: 
14167diff
changeset | 48 | val def_eq = Logic.mk_equals (Const(cname_full,ctype),choice_const ctype $ P) | 
| 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 skalberg parents: 
14167diff
changeset | 49 | val (thy',thms) = PureThy.add_defs_i covld [((cdefname,def_eq),[])] thy | 
| 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 skalberg parents: 
14167diff
changeset | 50 | val thm' = [thm,hd thms] MRS helper | 
| 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 skalberg parents: 
14167diff
changeset | 51 | in | 
| 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 skalberg parents: 
14167diff
changeset | 52 | mk_definitional cos (thy',thm') | 
| 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 skalberg parents: 
14167diff
changeset | 53 | end | 
| 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 skalberg parents: 
14167diff
changeset | 54 | 	  | _ => raise THM ("Internal error: Bad specification theorem",0,[thm])
 | 
| 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 skalberg parents: 
14167diff
changeset | 55 | |
| 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 skalberg parents: 
14167diff
changeset | 56 | fun mk_axiomatic axname cos arg = | 
| 14115 | 57 | let | 
| 14222 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 skalberg parents: 
14167diff
changeset | 58 | fun process [] (thy,tm) = | 
| 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 skalberg parents: 
14167diff
changeset | 59 | let | 
| 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 skalberg parents: 
14167diff
changeset | 60 | val (thy',thms) = PureThy.add_axioms_i [((axname,HOLogic.mk_Trueprop tm),[])] thy | 
| 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 skalberg parents: 
14167diff
changeset | 61 | in | 
| 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 skalberg parents: 
14167diff
changeset | 62 | (thy',hd thms) | 
| 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 skalberg parents: 
14167diff
changeset | 63 | end | 
| 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 skalberg parents: 
14167diff
changeset | 64 | | process ((thname,cname,covld)::cos) (thy,tm) = | 
| 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 skalberg parents: 
14167diff
changeset | 65 | case tm of | 
| 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 skalberg parents: 
14167diff
changeset | 66 | 		    Const("Ex",_) $ P =>
 | 
| 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 skalberg parents: 
14167diff
changeset | 67 | let | 
| 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 skalberg parents: 
14167diff
changeset | 68 | val ctype = domain_type (type_of P) | 
| 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 skalberg parents: 
14167diff
changeset | 69 | val cname_full = Sign.intern_const (sign_of thy) cname | 
| 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 skalberg parents: 
14167diff
changeset | 70 | val cdefname = if thname = "" | 
| 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 skalberg parents: 
14167diff
changeset | 71 | then Thm.def_name (Sign.base_name cname) | 
| 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 skalberg parents: 
14167diff
changeset | 72 | else thname | 
| 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 skalberg parents: 
14167diff
changeset | 73 | val co = Const(cname_full,ctype) | 
| 14223 
0ee05eef881b
Added support for making constants final, that is, ensuring that no
 skalberg parents: 
14222diff
changeset | 74 | val thy' = Theory.add_finals_i covld [co] thy | 
| 14222 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 skalberg parents: 
14167diff
changeset | 75 | val tm' = case P of | 
| 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 skalberg parents: 
14167diff
changeset | 76 | Abs(_, _, bodt) => subst_bound (co, bodt) | 
| 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 skalberg parents: 
14167diff
changeset | 77 | | _ => P $ co | 
| 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 skalberg parents: 
14167diff
changeset | 78 | in | 
| 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 skalberg parents: 
14167diff
changeset | 79 | process cos (thy',tm') | 
| 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 skalberg parents: 
14167diff
changeset | 80 | end | 
| 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 skalberg parents: 
14167diff
changeset | 81 | 		  | _ => raise TERM ("Internal error: Bad specification theorem",[tm])
 | 
| 14115 | 82 | in | 
| 14222 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 skalberg parents: 
14167diff
changeset | 83 | process cos arg | 
| 14115 | 84 | end | 
| 85 | ||
| 14222 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 skalberg parents: 
14167diff
changeset | 86 | in | 
| 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 skalberg parents: 
14167diff
changeset | 87 | fun proc_exprop axiomatic cos arg = | 
| 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 skalberg parents: 
14167diff
changeset | 88 | case axiomatic of | 
| 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 skalberg parents: 
14167diff
changeset | 89 | Some axname => mk_axiomatic axname cos (apsnd (HOLogic.dest_Trueprop o concl_of) arg) | 
| 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 skalberg parents: 
14167diff
changeset | 90 | | None => mk_definitional cos arg | 
| 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 skalberg parents: 
14167diff
changeset | 91 | end | 
| 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 skalberg parents: 
14167diff
changeset | 92 | |
| 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 skalberg parents: 
14167diff
changeset | 93 | fun add_specification_i axiomatic cos arg = | 
| 14121 
d2a0fd183f5f
Added handling of free variables (provided they are of sort HOL.type).
 skalberg parents: 
14118diff
changeset | 94 | arg |> apsnd freezeT | 
| 14222 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 skalberg parents: 
14167diff
changeset | 95 | |> proc_exprop axiomatic cos | 
| 14121 
d2a0fd183f5f
Added handling of free variables (provided they are of sort HOL.type).
 skalberg parents: 
14118diff
changeset | 96 | |> apsnd standard | 
| 14115 | 97 | |
| 98 | (* Collect all intances of constants in term *) | |
| 99 | ||
| 100 | fun collect_consts ( t $ u,tms) = collect_consts (u,collect_consts (t,tms)) | |
| 101 | | collect_consts ( Abs(_,_,t),tms) = collect_consts (t,tms) | |
| 102 | | collect_consts (tm as Const _,tms) = gen_ins (op aconv) (tm,tms) | |
| 103 | | collect_consts ( _,tms) = tms | |
| 104 | ||
| 105 | (* Complementing Type.varify... *) | |
| 106 | ||
| 107 | fun unvarify t fmap = | |
| 108 | let | |
| 109 | val fmap' = map Library.swap fmap | |
| 110 | fun unthaw (f as (a,S)) = | |
| 111 | (case assoc (fmap',a) of | |
| 112 | None => TVar f | |
| 113 | | Some b => TFree (b,S)) | |
| 114 | in | |
| 115 | map_term_types (map_type_tvar unthaw) t | |
| 116 | end | |
| 117 | ||
| 118 | (* The syntactic meddling needed to setup add_specification for work *) | |
| 119 | ||
| 14222 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 skalberg parents: 
14167diff
changeset | 120 | fun process_spec axiomatic cos alt_props int thy = | 
| 14115 | 121 | let | 
| 14167 | 122 | fun zip3 [] [] [] = [] | 
| 123 | | zip3 (x::xs) (y::ys) (z::zs) = (x,y,z)::zip3 xs ys zs | |
| 124 | | zip3 _ _ _ = error "SpecificationPackage.process_spec internal error" | |
| 125 | ||
| 126 | fun myfoldr f [x] = x | |
| 14164 
8c3fab596219
Allowed for splitting the specification over several lemmas.
 skalberg parents: 
14122diff
changeset | 127 | | myfoldr f (x::xs) = f (x,myfoldr f xs) | 
| 14167 | 128 | | myfoldr f [] = error "SpecificationPackage.process_spec internal error" | 
| 14164 
8c3fab596219
Allowed for splitting the specification over several lemmas.
 skalberg parents: 
14122diff
changeset | 129 | |
| 14115 | 130 | val sg = sign_of thy | 
| 14769 | 131 | fun typ_equiv t u = Sign.typ_instance sg (t,u) andalso Sign.typ_instance sg (u,t); | 
| 14164 
8c3fab596219
Allowed for splitting the specification over several lemmas.
 skalberg parents: 
14122diff
changeset | 132 | |
| 
8c3fab596219
Allowed for splitting the specification over several lemmas.
 skalberg parents: 
14122diff
changeset | 133 | val cprops = map (Thm.read_cterm sg o rpair Term.propT o snd) alt_props | 
| 
8c3fab596219
Allowed for splitting the specification over several lemmas.
 skalberg parents: 
14122diff
changeset | 134 | val ss = empty_ss setmksimps single addsimps [thm "HOL.atomize_imp",thm "HOL.atomize_all"] | 
| 
8c3fab596219
Allowed for splitting the specification over several lemmas.
 skalberg parents: 
14122diff
changeset | 135 | val rew_imps = map (Simplifier.full_rewrite ss) cprops | 
| 
8c3fab596219
Allowed for splitting the specification over several lemmas.
 skalberg parents: 
14122diff
changeset | 136 | val props' = map (HOLogic.dest_Trueprop o term_of o snd o dest_equals o cprop_of) rew_imps | 
| 14166 | 137 | |
| 138 | fun proc_single prop = | |
| 139 | let | |
| 140 | val frees = Term.term_frees prop | |
| 141 | val tsig = Sign.tsig_of (sign_of thy) | |
| 142 | val _ = assert (forall (fn v => Type.of_sort tsig (type_of v,HOLogic.typeS)) frees) | |
| 14167 | 143 | "Specificaton: Only free variables of sort 'type' allowed" | 
| 14166 | 144 | val prop_closed = foldr (fn ((vname,T),prop) => HOLogic.mk_all (vname,T,prop)) (map dest_Free frees,prop) | 
| 145 | in | |
| 146 | (prop_closed,frees) | |
| 147 | end | |
| 148 | ||
| 149 | val props'' = map proc_single props' | |
| 150 | val frees = map snd props'' | |
| 151 | val prop = myfoldr HOLogic.mk_conj (map fst props'') | |
| 14164 
8c3fab596219
Allowed for splitting the specification over several lemmas.
 skalberg parents: 
14122diff
changeset | 152 | val cprop = cterm_of sg (HOLogic.mk_Trueprop prop) | 
| 
8c3fab596219
Allowed for splitting the specification over several lemmas.
 skalberg parents: 
14122diff
changeset | 153 | |
| 14115 | 154 | val (prop_thawed,vmap) = Type.varify (prop,[]) | 
| 155 | val thawed_prop_consts = collect_consts (prop_thawed,[]) | |
| 14116 
63337d8e6e0f
Added optional theorem names for the constant definitions added during
 skalberg parents: 
14115diff
changeset | 156 | val (altcos,overloaded) = Library.split_list cos | 
| 
63337d8e6e0f
Added optional theorem names for the constant definitions added during
 skalberg parents: 
14115diff
changeset | 157 | val (names,sconsts) = Library.split_list altcos | 
| 14115 | 158 | val consts = map (term_of o Thm.read_cterm sg o rpair TypeInfer.logicT) sconsts | 
| 159 | val _ = assert (not (Library.exists (not o Term.is_Const) consts)) | |
| 14167 | 160 | "Specification: Non-constant found as parameter" | 
| 14166 | 161 | |
| 14115 | 162 | fun proc_const c = | 
| 163 | let | |
| 164 | val c' = fst (Type.varify (c,[])) | |
| 165 | val (cname,ctyp) = dest_Const c' | |
| 166 | in | |
| 167 | case filter (fn t => let val (name,typ) = dest_Const t | |
| 168 | in name = cname andalso typ_equiv typ ctyp | |
| 169 | end) thawed_prop_consts of | |
| 14167 | 170 | 		    [] => error ("Specification: No suitable instances of constant \"" ^ (Sign.string_of_term sg c) ^ "\" found")
 | 
| 14115 | 171 | | [cf] => unvarify cf vmap | 
| 14167 | 172 | 		  | _ => error ("Specification: Several variations of \"" ^ (Sign.string_of_term sg c) ^ "\" found (try applying explicit type constraints)")
 | 
| 14115 | 173 | end | 
| 174 | val proc_consts = map proc_const consts | |
| 175 | fun mk_exist (c,prop) = | |
| 176 | let | |
| 177 | val T = type_of c | |
| 14166 | 178 | val cname = Sign.base_name (fst (dest_Const c)) | 
| 179 | val vname = if Syntax.is_identifier cname | |
| 180 | then cname | |
| 181 | else "x" | |
| 14115 | 182 | in | 
| 14166 | 183 | HOLogic.exists_const T $ Abs(vname,T,Term.abstract_over (c,prop)) | 
| 14115 | 184 | end | 
| 14166 | 185 | val ex_prop = foldr mk_exist (proc_consts,prop) | 
| 14115 | 186 | val cnames = map (fst o dest_Const) proc_consts | 
| 14121 
d2a0fd183f5f
Added handling of free variables (provided they are of sort HOL.type).
 skalberg parents: 
14118diff
changeset | 187 | fun post_process (arg as (thy,thm)) = | 
| 
d2a0fd183f5f
Added handling of free variables (provided they are of sort HOL.type).
 skalberg parents: 
14118diff
changeset | 188 | let | 
| 
d2a0fd183f5f
Added handling of free variables (provided they are of sort HOL.type).
 skalberg parents: 
14118diff
changeset | 189 | fun inst_all sg (thm,v) = | 
| 
d2a0fd183f5f
Added handling of free variables (provided they are of sort HOL.type).
 skalberg parents: 
14118diff
changeset | 190 | let | 
| 
d2a0fd183f5f
Added handling of free variables (provided they are of sort HOL.type).
 skalberg parents: 
14118diff
changeset | 191 | val cv = cterm_of sg v | 
| 
d2a0fd183f5f
Added handling of free variables (provided they are of sort HOL.type).
 skalberg parents: 
14118diff
changeset | 192 | val cT = ctyp_of_term cv | 
| 
d2a0fd183f5f
Added handling of free variables (provided they are of sort HOL.type).
 skalberg parents: 
14118diff
changeset | 193 | val spec' = instantiate' [Some cT] [None,Some cv] spec | 
| 
d2a0fd183f5f
Added handling of free variables (provided they are of sort HOL.type).
 skalberg parents: 
14118diff
changeset | 194 | in | 
| 
d2a0fd183f5f
Added handling of free variables (provided they are of sort HOL.type).
 skalberg parents: 
14118diff
changeset | 195 | thm RS spec' | 
| 
d2a0fd183f5f
Added handling of free variables (provided they are of sort HOL.type).
 skalberg parents: 
14118diff
changeset | 196 | end | 
| 
d2a0fd183f5f
Added handling of free variables (provided they are of sort HOL.type).
 skalberg parents: 
14118diff
changeset | 197 | fun remove_alls frees thm = foldl (inst_all (sign_of_thm thm)) (thm,frees) | 
| 14167 | 198 | fun process_single ((name,atts),rew_imp,frees) args = | 
| 14164 
8c3fab596219
Allowed for splitting the specification over several lemmas.
 skalberg parents: 
14122diff
changeset | 199 | let | 
| 
8c3fab596219
Allowed for splitting the specification over several lemmas.
 skalberg parents: 
14122diff
changeset | 200 | fun undo_imps thm = | 
| 
8c3fab596219
Allowed for splitting the specification over several lemmas.
 skalberg parents: 
14122diff
changeset | 201 | equal_elim (symmetric rew_imp) thm | 
| 
8c3fab596219
Allowed for splitting the specification over several lemmas.
 skalberg parents: 
14122diff
changeset | 202 | |
| 
8c3fab596219
Allowed for splitting the specification over several lemmas.
 skalberg parents: 
14122diff
changeset | 203 | fun apply_atts arg = Thm.apply_attributes (arg,map (Attrib.global_attribute thy) atts) | 
| 
8c3fab596219
Allowed for splitting the specification over several lemmas.
 skalberg parents: 
14122diff
changeset | 204 | fun add_final (arg as (thy,thm)) = | 
| 
8c3fab596219
Allowed for splitting the specification over several lemmas.
 skalberg parents: 
14122diff
changeset | 205 | if name = "" | 
| 
8c3fab596219
Allowed for splitting the specification over several lemmas.
 skalberg parents: 
14122diff
changeset | 206 | then arg | 
| 
8c3fab596219
Allowed for splitting the specification over several lemmas.
 skalberg parents: 
14122diff
changeset | 207 | 			    else (writeln ("  " ^ name ^ ": " ^ (string_of_thm thm));
 | 
| 
8c3fab596219
Allowed for splitting the specification over several lemmas.
 skalberg parents: 
14122diff
changeset | 208 | PureThy.store_thm((name,thm),[]) thy) | 
| 
8c3fab596219
Allowed for splitting the specification over several lemmas.
 skalberg parents: 
14122diff
changeset | 209 | in | 
| 14166 | 210 | args |> apsnd (remove_alls frees) | 
| 14167 | 211 | |> apsnd undo_imps | 
| 14164 
8c3fab596219
Allowed for splitting the specification over several lemmas.
 skalberg parents: 
14122diff
changeset | 212 | |> apsnd standard | 
| 
8c3fab596219
Allowed for splitting the specification over several lemmas.
 skalberg parents: 
14122diff
changeset | 213 | |> apply_atts | 
| 
8c3fab596219
Allowed for splitting the specification over several lemmas.
 skalberg parents: 
14122diff
changeset | 214 | |> add_final | 
| 
8c3fab596219
Allowed for splitting the specification over several lemmas.
 skalberg parents: 
14122diff
changeset | 215 | end | 
| 
8c3fab596219
Allowed for splitting the specification over several lemmas.
 skalberg parents: 
14122diff
changeset | 216 | |
| 14167 | 217 | fun process_all [proc_arg] args = | 
| 218 | process_single proc_arg args | |
| 219 | | process_all (proc_arg::rest) (thy,thm) = | |
| 14164 
8c3fab596219
Allowed for splitting the specification over several lemmas.
 skalberg parents: 
14122diff
changeset | 220 | let | 
| 
8c3fab596219
Allowed for splitting the specification over several lemmas.
 skalberg parents: 
14122diff
changeset | 221 | val single_th = thm RS conjunct1 | 
| 
8c3fab596219
Allowed for splitting the specification over several lemmas.
 skalberg parents: 
14122diff
changeset | 222 | val rest_th = thm RS conjunct2 | 
| 14167 | 223 | val (thy',_) = process_single proc_arg (thy,single_th) | 
| 14164 
8c3fab596219
Allowed for splitting the specification over several lemmas.
 skalberg parents: 
14122diff
changeset | 224 | in | 
| 14167 | 225 | process_all rest (thy',rest_th) | 
| 14164 
8c3fab596219
Allowed for splitting the specification over several lemmas.
 skalberg parents: 
14122diff
changeset | 226 | end | 
| 14167 | 227 | | process_all [] _ = error "SpecificationPackage.process_spec internal error" | 
| 14164 
8c3fab596219
Allowed for splitting the specification over several lemmas.
 skalberg parents: 
14122diff
changeset | 228 | val alt_names = map fst alt_props | 
| 
8c3fab596219
Allowed for splitting the specification over several lemmas.
 skalberg parents: 
14122diff
changeset | 229 | val _ = if exists (fn(name,_) => not (name = "")) alt_names | 
| 
8c3fab596219
Allowed for splitting the specification over several lemmas.
 skalberg parents: 
14122diff
changeset | 230 | then writeln "specification" | 
| 
8c3fab596219
Allowed for splitting the specification over several lemmas.
 skalberg parents: 
14122diff
changeset | 231 | else () | 
| 14121 
d2a0fd183f5f
Added handling of free variables (provided they are of sort HOL.type).
 skalberg parents: 
14118diff
changeset | 232 | in | 
| 14122 
433f9a414537
Added handling of meta implication and meta quantification.
 skalberg parents: 
14121diff
changeset | 233 | arg |> apsnd freezeT | 
| 14167 | 234 | |> process_all (zip3 alt_names rew_imps frees) | 
| 14121 
d2a0fd183f5f
Added handling of free variables (provided they are of sort HOL.type).
 skalberg parents: 
14118diff
changeset | 235 | end | 
| 14115 | 236 | in | 
| 237 | IsarThy.theorem_i Drule.internalK | |
| 14222 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 skalberg parents: 
14167diff
changeset | 238 | 	    (("",[add_specification_i axiomatic (zip3 names cnames overloaded),post_process]),
 | 
| 14115 | 239 | (HOLogic.mk_Trueprop ex_prop,([],[]))) int thy | 
| 240 | end | |
| 241 | ||
| 242 | (* outer syntax *) | |
| 243 | ||
| 244 | local structure P = OuterParse and K = OuterSyntax.Keyword in | |
| 245 | ||
| 246 | (* taken from ~~/Pure/Isar/isar_syn.ML *) | |
| 247 | val opt_overloaded = | |
| 248 |   Scan.optional (P.$$$ "(" |-- P.!!! ((P.$$$ "overloaded" >> K true) --| P.$$$ ")")) false
 | |
| 249 | ||
| 14116 
63337d8e6e0f
Added optional theorem names for the constant definitions added during
 skalberg parents: 
14115diff
changeset | 250 | val opt_name = Scan.optional (P.name --| P.$$$ ":") "" | 
| 
63337d8e6e0f
Added optional theorem names for the constant definitions added during
 skalberg parents: 
14115diff
changeset | 251 | |
| 14115 | 252 | val specification_decl = | 
| 14116 
63337d8e6e0f
Added optional theorem names for the constant definitions added during
 skalberg parents: 
14115diff
changeset | 253 |     P.$$$ "(" |-- Scan.repeat1 (opt_name -- P.term -- opt_overloaded) --| P.$$$ ")" --
 | 
| 14164 
8c3fab596219
Allowed for splitting the specification over several lemmas.
 skalberg parents: 
14122diff
changeset | 254 | Scan.repeat1 (P.opt_thm_name ":" -- P.prop) | 
| 14115 | 255 | |
| 256 | val specificationP = | |
| 257 | OuterSyntax.command "specification" "define constants by specification" K.thy_goal | |
| 14164 
8c3fab596219
Allowed for splitting the specification over several lemmas.
 skalberg parents: 
14122diff
changeset | 258 | (specification_decl >> (fn (cos,alt_props) => | 
| 14222 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 skalberg parents: 
14167diff
changeset | 259 | Toplevel.print o (Toplevel.theory_to_proof | 
| 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 skalberg parents: 
14167diff
changeset | 260 | (process_spec None cos alt_props)))) | 
| 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 skalberg parents: 
14167diff
changeset | 261 | |
| 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 skalberg parents: 
14167diff
changeset | 262 | val ax_specification_decl = | 
| 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 skalberg parents: 
14167diff
changeset | 263 | P.name -- | 
| 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 skalberg parents: 
14167diff
changeset | 264 |     (P.$$$ "(" |-- Scan.repeat1 (opt_name -- P.term -- opt_overloaded) --| P.$$$ ")" --
 | 
| 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 skalberg parents: 
14167diff
changeset | 265 | Scan.repeat1 (P.opt_thm_name ":" -- P.prop)) | 
| 14115 | 266 | |
| 14222 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 skalberg parents: 
14167diff
changeset | 267 | val ax_specificationP = | 
| 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 skalberg parents: 
14167diff
changeset | 268 | OuterSyntax.command "ax_specification" "define constants by specification" K.thy_goal | 
| 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 skalberg parents: 
14167diff
changeset | 269 | (ax_specification_decl >> (fn (axname,(cos,alt_props)) => | 
| 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 skalberg parents: 
14167diff
changeset | 270 | Toplevel.print o (Toplevel.theory_to_proof | 
| 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 skalberg parents: 
14167diff
changeset | 271 | (process_spec (Some axname) cos alt_props)))) | 
| 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 skalberg parents: 
14167diff
changeset | 272 | |
| 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 skalberg parents: 
14167diff
changeset | 273 | val _ = OuterSyntax.add_parsers [specificationP,ax_specificationP] | 
| 14115 | 274 | |
| 275 | end | |
| 276 | ||
| 277 | ||
| 278 | end |