| author | wenzelm | 
| Tue, 03 Apr 2012 20:37:52 +0200 | |
| changeset 47319 | 8aa23a259ab2 | 
| parent 46974 | 7ca3608146d8 | 
| child 47815 | 43f677b3ae91 | 
| permissions | -rw-r--r-- | 
| 31723 
f5cafe803b55
discontinued ancient tradition to suffix certain ML module names with "_package"
 haftmann parents: 
30364diff
changeset | 1 | (* Title: HOL/Tools/choice_specification.ML | 
| 14115 | 2 | Author: Sebastian Skalberg, TU Muenchen | 
| 3 | ||
| 4 | Package for defining constants by specification. | |
| 5 | *) | |
| 6 | ||
| 31723 
f5cafe803b55
discontinued ancient tradition to suffix certain ML module names with "_package"
 haftmann parents: 
30364diff
changeset | 7 | signature CHOICE_SPECIFICATION = | 
| 14115 | 8 | sig | 
| 35807 
e4d1b5cbd429
added support for "specification" and "ax_specification" constructs to Nitpick
 blanchet parents: 
35625diff
changeset | 9 | val close_form : term -> term | 
| 18729 | 10 | val add_specification: string option -> (bstring * xstring * bool) list -> | 
| 11 | theory * thm -> theory * thm | |
| 14115 | 12 | end | 
| 13 | ||
| 31723 
f5cafe803b55
discontinued ancient tradition to suffix certain ML module names with "_package"
 haftmann parents: 
30364diff
changeset | 14 | structure Choice_Specification: CHOICE_SPECIFICATION = | 
| 14115 | 15 | struct | 
| 16 | ||
| 26478 | 17 | (* actual code *) | 
| 14115 | 18 | |
| 35807 
e4d1b5cbd429
added support for "specification" and "ax_specification" constructs to Nitpick
 blanchet parents: 
35625diff
changeset | 19 | fun close_form t = | 
| 
e4d1b5cbd429
added support for "specification" and "ax_specification" constructs to Nitpick
 blanchet parents: 
35625diff
changeset | 20 | fold_rev (fn (s, T) => fn t => HOLogic.mk_all (s, T, t)) | 
| 44121 | 21 | (map dest_Free (Misc_Legacy.term_frees t)) t | 
| 35807 
e4d1b5cbd429
added support for "specification" and "ax_specification" constructs to Nitpick
 blanchet parents: 
35625diff
changeset | 22 | |
| 46974 
7ca3608146d8
eliminated odd 'finalconsts' / Theory.add_finals;
 wenzelm parents: 
46961diff
changeset | 23 | fun add_final overloaded (c, T) thy = | 
| 
7ca3608146d8
eliminated odd 'finalconsts' / Theory.add_finals;
 wenzelm parents: 
46961diff
changeset | 24 | let | 
| 
7ca3608146d8
eliminated odd 'finalconsts' / Theory.add_finals;
 wenzelm parents: 
46961diff
changeset | 25 | val ctxt = Syntax.init_pretty_global thy; | 
| 
7ca3608146d8
eliminated odd 'finalconsts' / Theory.add_finals;
 wenzelm parents: 
46961diff
changeset | 26 | val _ = Theory.check_overloading ctxt overloaded (c, T); | 
| 
7ca3608146d8
eliminated odd 'finalconsts' / Theory.add_finals;
 wenzelm parents: 
46961diff
changeset | 27 | in Theory.add_deps ctxt "" (c, T) [] thy end; | 
| 
7ca3608146d8
eliminated odd 'finalconsts' / Theory.add_finals;
 wenzelm parents: 
46961diff
changeset | 28 | |
| 14222 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 skalberg parents: 
14167diff
changeset | 29 | local | 
| 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 skalberg parents: 
14167diff
changeset | 30 | fun mk_definitional [] arg = arg | 
| 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 skalberg parents: 
14167diff
changeset | 31 | | mk_definitional ((thname,cname,covld)::cos) (thy,thm) = | 
| 17336 | 32 | case HOLogic.dest_Trueprop (concl_of thm) of | 
| 38558 | 33 |             Const(@{const_name Ex},_) $ P =>
 | 
| 17336 | 34 | let | 
| 35 | val ctype = domain_type (type_of P) | |
| 22578 | 36 | val cname_full = Sign.intern_const thy cname | 
| 17336 | 37 | val cdefname = if thname = "" | 
| 30364 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30280diff
changeset | 38 | then Thm.def_name (Long_Name.base_name cname) | 
| 17336 | 39 | else thname | 
| 40 | val def_eq = Logic.mk_equals (Const(cname_full,ctype), | |
| 41 | HOLogic.choice_const ctype $ P) | |
| 39557 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 wenzelm parents: 
38756diff
changeset | 42 | val (thms, thy') = Global_Theory.add_defs covld [((Binding.name cdefname, def_eq),[])] thy | 
| 26478 | 43 |                 val thm' = [thm,hd thms] MRS @{thm exE_some}
 | 
| 17336 | 44 | in | 
| 45 | mk_definitional cos (thy',thm') | |
| 46 | end | |
| 47 |           | _ => raise THM ("Internal error: Bad specification theorem",0,[thm])
 | |
| 14222 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 skalberg parents: 
14167diff
changeset | 48 | |
| 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 skalberg parents: 
14167diff
changeset | 49 | fun mk_axiomatic axname cos arg = | 
| 17336 | 50 | let | 
| 51 | fun process [] (thy,tm) = | |
| 52 | let | |
| 35897 
8758895ea413
eliminated old-style Drule.add_axiom in favour of Specification.axiom, with explicit Drule.export_without_context to imitate the old behaviour;
 wenzelm parents: 
35856diff
changeset | 53 | val (thm, thy') = | 
| 
8758895ea413
eliminated old-style Drule.add_axiom in favour of Specification.axiom, with explicit Drule.export_without_context to imitate the old behaviour;
 wenzelm parents: 
35856diff
changeset | 54 | Specification.axiom ((Binding.name axname, []), HOLogic.mk_Trueprop tm) thy | 
| 17336 | 55 | in | 
| 35897 
8758895ea413
eliminated old-style Drule.add_axiom in favour of Specification.axiom, with explicit Drule.export_without_context to imitate the old behaviour;
 wenzelm parents: 
35856diff
changeset | 56 | (thy', Drule.export_without_context thm) | 
| 17336 | 57 | end | 
| 58 | | process ((thname,cname,covld)::cos) (thy,tm) = | |
| 59 | case tm of | |
| 38558 | 60 |                     Const(@{const_name Ex},_) $ P =>
 | 
| 17336 | 61 | let | 
| 62 | val ctype = domain_type (type_of P) | |
| 22578 | 63 | val cname_full = Sign.intern_const thy cname | 
| 17336 | 64 | val cdefname = if thname = "" | 
| 30364 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30280diff
changeset | 65 | then Thm.def_name (Long_Name.base_name cname) | 
| 17336 | 66 | else thname | 
| 46974 
7ca3608146d8
eliminated odd 'finalconsts' / Theory.add_finals;
 wenzelm parents: 
46961diff
changeset | 67 | val thy' = add_final covld (cname_full,ctype) thy | 
| 
7ca3608146d8
eliminated odd 'finalconsts' / Theory.add_finals;
 wenzelm parents: 
46961diff
changeset | 68 | val co = Const (cname_full,ctype) | 
| 17336 | 69 | val tm' = case P of | 
| 70 | Abs(_, _, bodt) => subst_bound (co, bodt) | |
| 71 | | _ => P $ co | |
| 72 | in | |
| 73 | process cos (thy',tm') | |
| 74 | end | |
| 75 |                   | _ => raise TERM ("Internal error: Bad specification theorem",[tm])
 | |
| 76 | in | |
| 77 | process cos arg | |
| 78 | end | |
| 14115 | 79 | |
| 14222 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 skalberg parents: 
14167diff
changeset | 80 | in | 
| 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 skalberg parents: 
14167diff
changeset | 81 | fun proc_exprop axiomatic cos arg = | 
| 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 skalberg parents: 
14167diff
changeset | 82 | case axiomatic of | 
| 17336 | 83 | SOME axname => mk_axiomatic axname cos (apsnd (HOLogic.dest_Trueprop o concl_of) arg) | 
| 15531 | 84 | | NONE => mk_definitional cos arg | 
| 14222 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 skalberg parents: 
14167diff
changeset | 85 | end | 
| 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 skalberg parents: 
14167diff
changeset | 86 | |
| 36618 
7a0990473e03
specification goal: eliminated redundant Thm.legacy_freezeT -- the goal is properly declared and should always produce fixed types in the result;
 wenzelm parents: 
36615diff
changeset | 87 | fun add_specification axiomatic cos = | 
| 
7a0990473e03
specification goal: eliminated redundant Thm.legacy_freezeT -- the goal is properly declared and should always produce fixed types in the result;
 wenzelm parents: 
36615diff
changeset | 88 | proc_exprop axiomatic cos | 
| 
7a0990473e03
specification goal: eliminated redundant Thm.legacy_freezeT -- the goal is properly declared and should always produce fixed types in the result;
 wenzelm parents: 
36615diff
changeset | 89 | #> apsnd Drule.export_without_context | 
| 14115 | 90 | |
| 18729 | 91 | |
| 14115 | 92 | (* Collect all intances of constants in term *) | 
| 93 | ||
| 94 | fun collect_consts ( t $ u,tms) = collect_consts (u,collect_consts (t,tms)) | |
| 95 | | collect_consts ( Abs(_,_,t),tms) = collect_consts (t,tms) | |
| 18921 | 96 | | collect_consts (tm as Const _,tms) = insert (op aconv) tm tms | 
| 14115 | 97 | | collect_consts ( _,tms) = tms | 
| 98 | ||
| 35845 
e5980f0ad025
renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
 wenzelm parents: 
35807diff
changeset | 99 | (* Complementing Type.varify_global... *) | 
| 14115 | 100 | |
| 35845 
e5980f0ad025
renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
 wenzelm parents: 
35807diff
changeset | 101 | fun unvarify_global t fmap = | 
| 14115 | 102 | let | 
| 17336 | 103 | val fmap' = map Library.swap fmap | 
| 17377 | 104 | fun unthaw (f as (a, S)) = | 
| 105 | (case AList.lookup (op =) fmap' a of | |
| 17336 | 106 | NONE => TVar f | 
| 17377 | 107 | | SOME (b, _) => TFree (b, S)) | 
| 14115 | 108 | in | 
| 20548 
8ef25fe585a8
renamed Term.map_term_types to Term.map_types (cf. Term.fold_types);
 wenzelm parents: 
20344diff
changeset | 109 | map_types (map_type_tvar unthaw) t | 
| 14115 | 110 | end | 
| 111 | ||
| 112 | (* The syntactic meddling needed to setup add_specification for work *) | |
| 113 | ||
| 17336 | 114 | fun process_spec axiomatic cos alt_props thy = | 
| 14115 | 115 | let | 
| 17336 | 116 | fun zip3 [] [] [] = [] | 
| 117 | | zip3 (x::xs) (y::ys) (z::zs) = (x,y,z)::zip3 xs ys zs | |
| 31723 
f5cafe803b55
discontinued ancient tradition to suffix certain ML module names with "_package"
 haftmann parents: 
30364diff
changeset | 118 | | zip3 _ _ _ = error "Choice_Specification.process_spec internal error" | 
| 14167 | 119 | |
| 17336 | 120 | fun myfoldr f [x] = x | 
| 121 | | myfoldr f (x::xs) = f (x,myfoldr f xs) | |
| 31723 
f5cafe803b55
discontinued ancient tradition to suffix certain ML module names with "_package"
 haftmann parents: 
30364diff
changeset | 122 | | myfoldr f [] = error "Choice_Specification.process_spec internal error" | 
| 14164 
8c3fab596219
Allowed for splitting the specification over several lemmas.
 skalberg parents: 
14122diff
changeset | 123 | |
| 17895 | 124 | val rew_imps = alt_props |> | 
| 35625 | 125 | map (Object_Logic.atomize o Thm.cterm_of thy o Syntax.read_prop_global thy o snd) | 
| 17895 | 126 | val props' = rew_imps |> | 
| 22902 
ac833b4bb7ee
moved some Drule operations to Thm (see more_thm.ML);
 wenzelm parents: 
22709diff
changeset | 127 | map (HOLogic.dest_Trueprop o term_of o snd o Thm.dest_equals o cprop_of) | 
| 14166 | 128 | |
| 17336 | 129 | fun proc_single prop = | 
| 130 | let | |
| 44121 | 131 | val frees = Misc_Legacy.term_frees prop | 
| 21858 
05f57309170c
avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
 wenzelm parents: 
21434diff
changeset | 132 | 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 | 133 | orelse error "Specificaton: Only free variables of sort 'type' allowed" | 
| 35807 
e4d1b5cbd429
added support for "specification" and "ax_specification" constructs to Nitpick
 blanchet parents: 
35625diff
changeset | 134 | val prop_closed = close_form prop | 
| 17336 | 135 | in | 
| 136 | (prop_closed,frees) | |
| 137 | end | |
| 14166 | 138 | |
| 17336 | 139 | val props'' = map proc_single props' | 
| 140 | val frees = map snd props'' | |
| 141 | val prop = myfoldr HOLogic.mk_conj (map fst props'') | |
| 17895 | 142 | val cprop = cterm_of thy (HOLogic.mk_Trueprop prop) | 
| 14164 
8c3fab596219
Allowed for splitting the specification over several lemmas.
 skalberg parents: 
14122diff
changeset | 143 | |
| 35845 
e5980f0ad025
renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
 wenzelm parents: 
35807diff
changeset | 144 | val (vmap, prop_thawed) = Type.varify_global [] prop | 
| 17336 | 145 | val thawed_prop_consts = collect_consts (prop_thawed,[]) | 
| 146 | val (altcos,overloaded) = Library.split_list cos | |
| 147 | val (names,sconsts) = Library.split_list altcos | |
| 24707 | 148 | 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 | 149 | 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 | 150 | orelse error "Specification: Non-constant found as parameter" | 
| 14166 | 151 | |
| 17336 | 152 | fun proc_const c = | 
| 153 | let | |
| 35845 
e5980f0ad025
renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
 wenzelm parents: 
35807diff
changeset | 154 | val (_, c') = Type.varify_global [] c | 
| 17336 | 155 | val (cname,ctyp) = dest_Const c' | 
| 156 | in | |
| 33317 | 157 | case filter (fn t => let val (name,typ) = dest_Const t | 
| 19414 | 158 | in name = cname andalso Sign.typ_equiv thy (typ, ctyp) | 
| 17336 | 159 | end) thawed_prop_consts of | 
| 26939 
1035c89b4c02
moved global pretty/string_of functions from Sign to Syntax;
 wenzelm parents: 
26928diff
changeset | 160 |                     [] => error ("Specification: No suitable instances of constant \"" ^ Syntax.string_of_term_global thy c ^ "\" found")
 | 
| 35845 
e5980f0ad025
renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
 wenzelm parents: 
35807diff
changeset | 161 | | [cf] => unvarify_global cf vmap | 
| 26939 
1035c89b4c02
moved global pretty/string_of functions from Sign to Syntax;
 wenzelm parents: 
26928diff
changeset | 162 |                   | _ => error ("Specification: Several variations of \"" ^ Syntax.string_of_term_global thy c ^ "\" found (try applying explicit type constraints)")
 | 
| 17336 | 163 | end | 
| 164 | val proc_consts = map proc_const consts | |
| 33339 | 165 | fun mk_exist c prop = | 
| 17336 | 166 | let | 
| 167 | val T = type_of c | |
| 30364 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30280diff
changeset | 168 | val cname = Long_Name.base_name (fst (dest_Const c)) | 
| 42290 
b1f544c84040
discontinued special treatment of structure Lexicon;
 wenzelm parents: 
39557diff
changeset | 169 | val vname = if Lexicon.is_identifier cname | 
| 17336 | 170 | then cname | 
| 171 | else "x" | |
| 172 | in | |
| 173 | HOLogic.exists_const T $ Abs(vname,T,Term.abstract_over (c,prop)) | |
| 174 | end | |
| 33339 | 175 | val ex_prop = fold_rev mk_exist proc_consts prop | 
| 17336 | 176 | val cnames = map (fst o dest_Const) proc_consts | 
| 177 | fun post_process (arg as (thy,thm)) = | |
| 178 | let | |
| 33245 | 179 | fun inst_all thy v thm = | 
| 17336 | 180 | let | 
| 17895 | 181 | val cv = cterm_of thy v | 
| 17336 | 182 | val cT = ctyp_of_term cv | 
| 183 | val spec' = instantiate' [SOME cT] [NONE,SOME cv] spec | |
| 184 | in | |
| 185 | thm RS spec' | |
| 186 | end | |
| 22578 | 187 | fun remove_alls frees thm = | 
| 33245 | 188 | fold (inst_all (Thm.theory_of_thm thm)) frees thm | 
| 17336 | 189 | fun process_single ((name,atts),rew_imp,frees) args = | 
| 190 | let | |
| 191 | fun undo_imps thm = | |
| 36945 | 192 | Thm.equal_elim (Thm.symmetric rew_imp) thm | 
| 14164 
8c3fab596219
Allowed for splitting the specification over several lemmas.
 skalberg parents: 
14122diff
changeset | 193 | |
| 46775 
6287653e63ec
canonical argument order for attribute application;
 wenzelm parents: 
45592diff
changeset | 194 | fun add_final (thm, thy) = | 
| 17336 | 195 | if name = "" | 
| 46775 
6287653e63ec
canonical argument order for attribute application;
 wenzelm parents: 
45592diff
changeset | 196 | then (thm, thy) | 
| 32091 
30e2ffbba718
proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
 wenzelm parents: 
31723diff
changeset | 197 |                             else (writeln ("  " ^ name ^ ": " ^ Display.string_of_thm_global thy thm);
 | 
| 39557 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 wenzelm parents: 
38756diff
changeset | 198 | Global_Theory.store_thm (Binding.name name, thm) thy) | 
| 17336 | 199 | in | 
| 46775 
6287653e63ec
canonical argument order for attribute application;
 wenzelm parents: 
45592diff
changeset | 200 | swap args | 
| 
6287653e63ec
canonical argument order for attribute application;
 wenzelm parents: 
45592diff
changeset | 201 | |> apfst (remove_alls frees) | 
| 
6287653e63ec
canonical argument order for attribute application;
 wenzelm parents: 
45592diff
changeset | 202 | |> apfst undo_imps | 
| 
6287653e63ec
canonical argument order for attribute application;
 wenzelm parents: 
45592diff
changeset | 203 | |> apfst Drule.export_without_context | 
| 
6287653e63ec
canonical argument order for attribute application;
 wenzelm parents: 
45592diff
changeset | 204 | |-> Thm.theory_attributes | 
| 45592 | 205 | (map (Attrib.attribute thy) | 
| 206 |                                   (@{attributes [nitpick_choice_spec]} @ atts))
 | |
| 17336 | 207 | |> add_final | 
| 46775 
6287653e63ec
canonical argument order for attribute application;
 wenzelm parents: 
45592diff
changeset | 208 | |> swap | 
| 17336 | 209 | end | 
| 14164 
8c3fab596219
Allowed for splitting the specification over several lemmas.
 skalberg parents: 
14122diff
changeset | 210 | |
| 17336 | 211 | fun process_all [proc_arg] args = | 
| 212 | process_single proc_arg args | |
| 213 | | process_all (proc_arg::rest) (thy,thm) = | |
| 214 | let | |
| 215 | val single_th = thm RS conjunct1 | |
| 216 | val rest_th = thm RS conjunct2 | |
| 217 | val (thy',_) = process_single proc_arg (thy,single_th) | |
| 218 | in | |
| 219 | process_all rest (thy',rest_th) | |
| 220 | end | |
| 31723 
f5cafe803b55
discontinued ancient tradition to suffix certain ML module names with "_package"
 haftmann parents: 
30364diff
changeset | 221 | | process_all [] _ = error "Choice_Specification.process_spec internal error" | 
| 17336 | 222 | val alt_names = map fst alt_props | 
| 223 | val _ = if exists (fn(name,_) => not (name = "")) alt_names | |
| 224 | then writeln "specification" | |
| 225 | else () | |
| 226 | in | |
| 36618 
7a0990473e03
specification goal: eliminated redundant Thm.legacy_freezeT -- the goal is properly declared and should always produce fixed types in the result;
 wenzelm parents: 
36615diff
changeset | 227 | arg |> apsnd Thm.unvarify_global | 
| 17336 | 228 | |> process_all (zip3 alt_names rew_imps frees) | 
| 229 | end | |
| 20903 | 230 | |
| 42361 | 231 | fun after_qed [[thm]] = Proof_Context.background_theory (fn thy => | 
| 38756 
d07959fabde6
renamed ProofContext.theory(_result) to ProofContext.background_theory(_result) to emphasize that this belongs to the infrastructure and is rarely appropriate in user-space tools;
 wenzelm parents: 
38558diff
changeset | 232 | #1 (post_process (add_specification axiomatic (zip3 names cnames overloaded) (thy, thm)))); | 
| 14115 | 233 | in | 
| 20903 | 234 | thy | 
| 42361 | 235 | |> Proof_Context.init_global | 
| 36618 
7a0990473e03
specification goal: eliminated redundant Thm.legacy_freezeT -- the goal is properly declared and should always produce fixed types in the result;
 wenzelm parents: 
36615diff
changeset | 236 | |> Variable.declare_term ex_prop | 
| 36323 
655e2d74de3a
modernized naming conventions of main Isar proof elements;
 wenzelm parents: 
35897diff
changeset | 237 | |> Proof.theorem NONE after_qed [[(HOLogic.mk_Trueprop ex_prop, [])]] | 
| 20903 | 238 | end; | 
| 14115 | 239 | |
| 17336 | 240 | |
| 14115 | 241 | (* outer syntax *) | 
| 242 | ||
| 46949 | 243 | val opt_name = Scan.optional (Parse.name --| @{keyword ":"}) ""
 | 
| 36960 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
36954diff
changeset | 244 | val opt_overloaded = Parse.opt_keyword "overloaded" | 
| 14116 
63337d8e6e0f
Added optional theorem names for the constant definitions added during
 skalberg parents: 
14115diff
changeset | 245 | |
| 46961 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 wenzelm parents: 
46949diff
changeset | 246 | val _ = | 
| 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 wenzelm parents: 
46949diff
changeset | 247 |   Outer_Syntax.command @{command_spec "specification"} "define constants by specification"
 | 
| 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 wenzelm parents: 
46949diff
changeset | 248 |     (@{keyword "("} |-- Scan.repeat1 (opt_name -- Parse.term -- opt_overloaded) --| @{keyword ")"} --
 | 
| 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 wenzelm parents: 
46949diff
changeset | 249 | Scan.repeat1 ((Parse_Spec.opt_thm_name ":" >> apfst Binding.name_of) -- Parse.prop) | 
| 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 wenzelm parents: 
46949diff
changeset | 250 | >> (fn (cos, alt_props) => Toplevel.print o | 
| 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 wenzelm parents: 
46949diff
changeset | 251 | (Toplevel.theory_to_proof (process_spec NONE cos alt_props)))) | 
| 14115 | 252 | |
| 24867 | 253 | val _ = | 
| 46961 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 wenzelm parents: 
46949diff
changeset | 254 |   Outer_Syntax.command @{command_spec "ax_specification"} "define constants by specification"
 | 
| 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 wenzelm parents: 
46949diff
changeset | 255 | (Parse.name -- | 
| 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 wenzelm parents: 
46949diff
changeset | 256 |       (@{keyword "("} |-- Scan.repeat1 (opt_name -- Parse.term -- opt_overloaded) --| @{keyword ")"} --
 | 
| 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 wenzelm parents: 
46949diff
changeset | 257 | Scan.repeat1 ((Parse_Spec.opt_thm_name ":" >> apfst Binding.name_of) -- Parse.prop)) | 
| 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 wenzelm parents: 
46949diff
changeset | 258 | >> (fn (axname, (cos, alt_props)) => | 
| 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 wenzelm parents: 
46949diff
changeset | 259 | Toplevel.print o (Toplevel.theory_to_proof (process_spec (SOME axname) cos alt_props)))) | 
| 14222 
1e61f23fd359
Added axiomatic specifications (ax_specification).
 skalberg parents: 
14167diff
changeset | 260 | |
| 14115 | 261 | end |