primitive definitions of bot/top/inf/sup for bool and fun are named with canonical suffix `_def` rather than `_eq`
1 (* Title: HOL/Tools/choice_specification.ML
2 Author: Sebastian Skalberg, TU Muenchen
4 Package for defining constants by specification.
7 signature CHOICE_SPECIFICATION =
9 val close_form : term -> term
10 val add_specification: string option -> (bstring * xstring * bool) list ->
11 theory * thm -> theory * thm
14 structure Choice_Specification: CHOICE_SPECIFICATION =
20 fold_rev (fn (s, T) => fn t => HOLogic.mk_all (s, T, t))
21 (map dest_Free (OldTerm.term_frees t)) t
24 fun mk_definitional [] arg = arg
25 | mk_definitional ((thname,cname,covld)::cos) (thy,thm) =
26 case HOLogic.dest_Trueprop (concl_of thm) of
27 Const(@{const_name Ex},_) $ P =>
29 val ctype = domain_type (type_of P)
30 val cname_full = Sign.intern_const thy cname
31 val cdefname = if thname = ""
32 then Thm.def_name (Long_Name.base_name cname)
34 val def_eq = Logic.mk_equals (Const(cname_full,ctype),
35 HOLogic.choice_const ctype $ P)
36 val (thms, thy') = Global_Theory.add_defs covld [((Binding.name cdefname, def_eq),[])] thy
37 val thm' = [thm,hd thms] MRS @{thm exE_some}
39 mk_definitional cos (thy',thm')
41 | _ => raise THM ("Internal error: Bad specification theorem",0,[thm])
43 fun mk_axiomatic axname cos arg =
45 fun process [] (thy,tm) =
48 Specification.axiom ((Binding.name axname, []), HOLogic.mk_Trueprop tm) thy
50 (thy', Drule.export_without_context thm)
52 | process ((thname,cname,covld)::cos) (thy,tm) =
54 Const(@{const_name Ex},_) $ P =>
56 val ctype = domain_type (type_of P)
57 val cname_full = Sign.intern_const thy cname
58 val cdefname = if thname = ""
59 then Thm.def_name (Long_Name.base_name cname)
61 val co = Const(cname_full,ctype)
62 val thy' = Theory.add_finals_i covld [co] thy
64 Abs(_, _, bodt) => subst_bound (co, bodt)
67 process cos (thy',tm')
69 | _ => raise TERM ("Internal error: Bad specification theorem",[tm])
75 fun proc_exprop axiomatic cos arg =
77 SOME axname => mk_axiomatic axname cos (apsnd (HOLogic.dest_Trueprop o concl_of) arg)
78 | NONE => mk_definitional cos arg
81 fun add_specification axiomatic cos =
82 proc_exprop axiomatic cos
83 #> apsnd Drule.export_without_context
86 (* Collect all intances of constants in term *)
88 fun collect_consts ( t $ u,tms) = collect_consts (u,collect_consts (t,tms))
89 | collect_consts ( Abs(_,_,t),tms) = collect_consts (t,tms)
90 | collect_consts (tm as Const _,tms) = insert (op aconv) tm tms
91 | collect_consts ( _,tms) = tms
93 (* Complementing Type.varify_global... *)
95 fun unvarify_global t fmap =
97 val fmap' = map Library.swap fmap
98 fun unthaw (f as (a, S)) =
99 (case AList.lookup (op =) fmap' a of
101 | SOME (b, _) => TFree (b, S))
103 map_types (map_type_tvar unthaw) t
106 (* The syntactic meddling needed to setup add_specification for work *)
108 fun process_spec axiomatic cos alt_props thy =
110 fun zip3 [] [] [] = []
111 | zip3 (x::xs) (y::ys) (z::zs) = (x,y,z)::zip3 xs ys zs
112 | zip3 _ _ _ = error "Choice_Specification.process_spec internal error"
114 fun myfoldr f [x] = x
115 | myfoldr f (x::xs) = f (x,myfoldr f xs)
116 | myfoldr f [] = error "Choice_Specification.process_spec internal error"
118 val rew_imps = alt_props |>
119 map (Object_Logic.atomize o Thm.cterm_of thy o Syntax.read_prop_global thy o snd)
120 val props' = rew_imps |>
121 map (HOLogic.dest_Trueprop o term_of o snd o Thm.dest_equals o cprop_of)
123 fun proc_single prop =
125 val frees = OldTerm.term_frees prop
126 val _ = forall (fn v => Sign.of_sort thy (type_of v,HOLogic.typeS)) frees
127 orelse error "Specificaton: Only free variables of sort 'type' allowed"
128 val prop_closed = close_form prop
133 val props'' = map proc_single props'
134 val frees = map snd props''
135 val prop = myfoldr HOLogic.mk_conj (map fst props'')
136 val cprop = cterm_of thy (HOLogic.mk_Trueprop prop)
138 val (vmap, prop_thawed) = Type.varify_global [] prop
139 val thawed_prop_consts = collect_consts (prop_thawed,[])
140 val (altcos,overloaded) = Library.split_list cos
141 val (names,sconsts) = Library.split_list altcos
142 val consts = map (Syntax.read_term_global thy) sconsts
143 val _ = not (Library.exists (not o Term.is_Const) consts)
144 orelse error "Specification: Non-constant found as parameter"
148 val (_, c') = Type.varify_global [] c
149 val (cname,ctyp) = dest_Const c'
151 case filter (fn t => let val (name,typ) = dest_Const t
152 in name = cname andalso Sign.typ_equiv thy (typ, ctyp)
153 end) thawed_prop_consts of
154 [] => error ("Specification: No suitable instances of constant \"" ^ Syntax.string_of_term_global thy c ^ "\" found")
155 | [cf] => unvarify_global cf vmap
156 | _ => error ("Specification: Several variations of \"" ^ Syntax.string_of_term_global thy c ^ "\" found (try applying explicit type constraints)")
158 val proc_consts = map proc_const consts
159 fun mk_exist c prop =
162 val cname = Long_Name.base_name (fst (dest_Const c))
163 val vname = if Syntax.is_identifier cname
167 HOLogic.exists_const T $ Abs(vname,T,Term.abstract_over (c,prop))
169 val ex_prop = fold_rev mk_exist proc_consts prop
170 val cnames = map (fst o dest_Const) proc_consts
171 fun post_process (arg as (thy,thm)) =
173 fun inst_all thy v thm =
175 val cv = cterm_of thy v
176 val cT = ctyp_of_term cv
177 val spec' = instantiate' [SOME cT] [NONE,SOME cv] spec
181 fun remove_alls frees thm =
182 fold (inst_all (Thm.theory_of_thm thm)) frees thm
183 fun process_single ((name,atts),rew_imp,frees) args =
186 Thm.equal_elim (Thm.symmetric rew_imp) thm
188 fun add_final (arg as (thy, thm)) =
190 then arg |> Library.swap
191 else (writeln (" " ^ name ^ ": " ^ Display.string_of_thm_global thy thm);
192 Global_Theory.store_thm (Binding.name name, thm) thy)
194 args |> apsnd (remove_alls frees)
196 |> apsnd Drule.export_without_context
197 |> Thm.theory_attributes (map (Attrib.attribute thy)
198 (Attrib.internal (K Nitpick_Choice_Specs.add) :: atts))
203 fun process_all [proc_arg] args =
204 process_single proc_arg args
205 | process_all (proc_arg::rest) (thy,thm) =
207 val single_th = thm RS conjunct1
208 val rest_th = thm RS conjunct2
209 val (thy',_) = process_single proc_arg (thy,single_th)
211 process_all rest (thy',rest_th)
213 | process_all [] _ = error "Choice_Specification.process_spec internal error"
214 val alt_names = map fst alt_props
215 val _ = if exists (fn(name,_) => not (name = "")) alt_names
216 then writeln "specification"
219 arg |> apsnd Thm.unvarify_global
220 |> process_all (zip3 alt_names rew_imps frees)
223 fun after_qed [[thm]] = ProofContext.background_theory (fn thy =>
224 #1 (post_process (add_specification axiomatic (zip3 names cnames overloaded) (thy, thm))));
227 |> ProofContext.init_global
228 |> Variable.declare_term ex_prop
229 |> Proof.theorem NONE after_qed [[(HOLogic.mk_Trueprop ex_prop, [])]]
235 val opt_name = Scan.optional (Parse.name --| Parse.$$$ ":") ""
236 val opt_overloaded = Parse.opt_keyword "overloaded"
238 val specification_decl =
239 Parse.$$$ "(" |-- Scan.repeat1 (opt_name -- Parse.term -- opt_overloaded) --| Parse.$$$ ")" --
240 Scan.repeat1 ((Parse_Spec.opt_thm_name ":" >> apfst Binding.name_of) -- Parse.prop)
243 Outer_Syntax.command "specification" "define constants by specification" Keyword.thy_goal
244 (specification_decl >> (fn (cos,alt_props) =>
245 Toplevel.print o (Toplevel.theory_to_proof
246 (process_spec NONE cos alt_props))))
248 val ax_specification_decl =
250 (Parse.$$$ "(" |-- Scan.repeat1 (opt_name -- Parse.term -- opt_overloaded) --| Parse.$$$ ")" --
251 Scan.repeat1 ((Parse_Spec.opt_thm_name ":" >> apfst Binding.name_of) -- Parse.prop))
254 Outer_Syntax.command "ax_specification" "define constants by specification" Keyword.thy_goal
255 (ax_specification_decl >> (fn (axname,(cos,alt_props)) =>
256 Toplevel.print o (Toplevel.theory_to_proof
257 (process_spec (SOME axname) cos alt_props))))