skalberg@14115: (* Title: HOL/Tools/specification_package.ML skalberg@14115: ID: $Id$ skalberg@14115: Author: Sebastian Skalberg, TU Muenchen skalberg@14115: skalberg@14115: Package for defining constants by specification. skalberg@14115: *) skalberg@14115: skalberg@14115: signature SPECIFICATION_PACKAGE = skalberg@14115: sig wenzelm@17336: val quiet_mode: bool ref wenzelm@18729: val add_specification: string option -> (bstring * xstring * bool) list -> wenzelm@18729: theory * thm -> theory * thm skalberg@14115: end skalberg@14115: wenzelm@17895: structure SpecificationPackage: SPECIFICATION_PACKAGE = skalberg@14115: struct skalberg@14115: skalberg@14115: (* messages *) skalberg@14115: skalberg@14115: val quiet_mode = ref false skalberg@14115: fun message s = if ! quiet_mode then () else writeln s skalberg@14115: skalberg@14115: skalberg@14115: (* Actual code *) skalberg@14115: skalberg@14222: local wenzelm@17895: val exE_some = thm "exE_some"; wenzelm@17895: skalberg@14222: fun mk_definitional [] arg = arg skalberg@14222: | mk_definitional ((thname,cname,covld)::cos) (thy,thm) = wenzelm@17336: case HOLogic.dest_Trueprop (concl_of thm) of wenzelm@17336: Const("Ex",_) $ P => wenzelm@17336: let wenzelm@17336: val ctype = domain_type (type_of P) wenzelm@17336: val cname_full = Sign.intern_const (sign_of thy) cname wenzelm@17336: val cdefname = if thname = "" wenzelm@17336: then Thm.def_name (Sign.base_name cname) wenzelm@17336: else thname wenzelm@17336: val def_eq = Logic.mk_equals (Const(cname_full,ctype), wenzelm@17336: HOLogic.choice_const ctype $ P) haftmann@18358: val (thms, thy') = PureThy.add_defs_i covld [((cdefname,def_eq),[])] thy wenzelm@17895: val thm' = [thm,hd thms] MRS exE_some wenzelm@17336: in wenzelm@17336: mk_definitional cos (thy',thm') wenzelm@17336: end wenzelm@17336: | _ => raise THM ("Internal error: Bad specification theorem",0,[thm]) skalberg@14222: skalberg@14222: fun mk_axiomatic axname cos arg = wenzelm@17336: let wenzelm@17336: fun process [] (thy,tm) = wenzelm@17336: let haftmann@18377: val (thms, thy') = PureThy.add_axioms_i [((axname,HOLogic.mk_Trueprop tm),[])] thy wenzelm@17336: in wenzelm@17336: (thy',hd thms) wenzelm@17336: end wenzelm@17336: | process ((thname,cname,covld)::cos) (thy,tm) = wenzelm@17336: case tm of wenzelm@17336: Const("Ex",_) $ P => wenzelm@17336: let wenzelm@17336: val ctype = domain_type (type_of P) wenzelm@17336: val cname_full = Sign.intern_const (sign_of thy) cname wenzelm@17336: val cdefname = if thname = "" wenzelm@17336: then Thm.def_name (Sign.base_name cname) wenzelm@17336: else thname wenzelm@17336: val co = Const(cname_full,ctype) wenzelm@17336: val thy' = Theory.add_finals_i covld [co] thy wenzelm@17336: val tm' = case P of wenzelm@17336: Abs(_, _, bodt) => subst_bound (co, bodt) wenzelm@17336: | _ => P $ co wenzelm@17336: in wenzelm@17336: process cos (thy',tm') wenzelm@17336: end wenzelm@17336: | _ => raise TERM ("Internal error: Bad specification theorem",[tm]) wenzelm@17336: in wenzelm@17336: process cos arg wenzelm@17336: end skalberg@14115: skalberg@14222: in skalberg@14222: fun proc_exprop axiomatic cos arg = skalberg@14222: case axiomatic of wenzelm@17336: SOME axname => mk_axiomatic axname cos (apsnd (HOLogic.dest_Trueprop o concl_of) arg) skalberg@15531: | NONE => mk_definitional cos arg skalberg@14222: end skalberg@14222: wenzelm@18729: fun add_specification axiomatic cos arg = wenzelm@19876: arg |> apsnd Thm.freezeT wenzelm@17336: |> proc_exprop axiomatic cos wenzelm@17336: |> apsnd standard skalberg@14115: wenzelm@18729: fun add_spec x y (context, thm) = wenzelm@18729: (Context.the_theory context, thm) |> add_specification x y |>> Context.Theory; wenzelm@18729: wenzelm@18729: skalberg@14115: (* Collect all intances of constants in term *) skalberg@14115: skalberg@14115: fun collect_consts ( t $ u,tms) = collect_consts (u,collect_consts (t,tms)) skalberg@14115: | collect_consts ( Abs(_,_,t),tms) = collect_consts (t,tms) wenzelm@18921: | collect_consts (tm as Const _,tms) = insert (op aconv) tm tms skalberg@14115: | collect_consts ( _,tms) = tms skalberg@14115: skalberg@14115: (* Complementing Type.varify... *) skalberg@14115: skalberg@14115: fun unvarify t fmap = skalberg@14115: let wenzelm@17336: val fmap' = map Library.swap fmap haftmann@17377: fun unthaw (f as (a, S)) = haftmann@17377: (case AList.lookup (op =) fmap' a of wenzelm@17336: NONE => TVar f haftmann@17377: | SOME (b, _) => TFree (b, S)) skalberg@14115: in wenzelm@20548: map_types (map_type_tvar unthaw) t skalberg@14115: end skalberg@14115: skalberg@14115: (* The syntactic meddling needed to setup add_specification for work *) skalberg@14115: wenzelm@17336: fun process_spec axiomatic cos alt_props thy = skalberg@14115: let wenzelm@17336: fun zip3 [] [] [] = [] wenzelm@17336: | zip3 (x::xs) (y::ys) (z::zs) = (x,y,z)::zip3 xs ys zs wenzelm@17336: | zip3 _ _ _ = error "SpecificationPackage.process_spec internal error" skalberg@14167: wenzelm@17336: fun myfoldr f [x] = x wenzelm@17336: | myfoldr f (x::xs) = f (x,myfoldr f xs) wenzelm@17336: | myfoldr f [] = error "SpecificationPackage.process_spec internal error" skalberg@14164: wenzelm@17895: val rew_imps = alt_props |> wenzelm@18776: map (ObjectLogic.atomize_cterm o Thm.read_cterm thy o rpair Term.propT o snd) wenzelm@17895: val props' = rew_imps |> wenzelm@17895: map (HOLogic.dest_Trueprop o term_of o snd o dest_equals o cprop_of) skalberg@14166: wenzelm@17336: fun proc_single prop = wenzelm@17336: let wenzelm@17336: val frees = Term.term_frees prop wenzelm@20344: val _ = assert (forall (fn v => Sign.of_sort thy (type_of v,HOLogic.typeS)) frees) wenzelm@17336: "Specificaton: Only free variables of sort 'type' allowed" wenzelm@17336: val prop_closed = foldr (fn ((vname,T),prop) => HOLogic.mk_all (vname,T,prop)) prop (map dest_Free frees) wenzelm@17336: in wenzelm@17336: (prop_closed,frees) wenzelm@17336: end skalberg@14166: wenzelm@17336: val props'' = map proc_single props' wenzelm@17336: val frees = map snd props'' wenzelm@17336: val prop = myfoldr HOLogic.mk_conj (map fst props'') wenzelm@17895: val cprop = cterm_of thy (HOLogic.mk_Trueprop prop) skalberg@14164: wenzelm@17336: val (prop_thawed,vmap) = Type.varify (prop,[]) wenzelm@17336: val thawed_prop_consts = collect_consts (prop_thawed,[]) wenzelm@17336: val (altcos,overloaded) = Library.split_list cos wenzelm@17336: val (names,sconsts) = Library.split_list altcos wenzelm@17895: val consts = map (term_of o Thm.read_cterm thy o rpair TypeInfer.logicT) sconsts wenzelm@17336: val _ = assert (not (Library.exists (not o Term.is_Const) consts)) wenzelm@17336: "Specification: Non-constant found as parameter" skalberg@14166: wenzelm@17336: fun proc_const c = wenzelm@17336: let wenzelm@17336: val c' = fst (Type.varify (c,[])) wenzelm@17336: val (cname,ctyp) = dest_Const c' wenzelm@17336: in wenzelm@17336: case List.filter (fn t => let val (name,typ) = dest_Const t wenzelm@19414: in name = cname andalso Sign.typ_equiv thy (typ, ctyp) wenzelm@17336: end) thawed_prop_consts of wenzelm@17895: [] => error ("Specification: No suitable instances of constant \"" ^ (Sign.string_of_term thy c) ^ "\" found") wenzelm@17336: | [cf] => unvarify cf vmap wenzelm@17895: | _ => error ("Specification: Several variations of \"" ^ (Sign.string_of_term thy c) ^ "\" found (try applying explicit type constraints)") wenzelm@17336: end wenzelm@17336: val proc_consts = map proc_const consts wenzelm@17336: fun mk_exist (c,prop) = wenzelm@17336: let wenzelm@17336: val T = type_of c wenzelm@17336: val cname = Sign.base_name (fst (dest_Const c)) wenzelm@17336: val vname = if Syntax.is_identifier cname wenzelm@17336: then cname wenzelm@17336: else "x" wenzelm@17336: in wenzelm@17336: HOLogic.exists_const T $ Abs(vname,T,Term.abstract_over (c,prop)) wenzelm@17336: end wenzelm@17336: val ex_prop = foldr mk_exist prop proc_consts wenzelm@17336: val cnames = map (fst o dest_Const) proc_consts wenzelm@17336: fun post_process (arg as (thy,thm)) = wenzelm@17336: let wenzelm@17895: fun inst_all thy (thm,v) = wenzelm@17336: let wenzelm@17895: val cv = cterm_of thy v wenzelm@17336: val cT = ctyp_of_term cv wenzelm@17336: val spec' = instantiate' [SOME cT] [NONE,SOME cv] spec wenzelm@17336: in wenzelm@17336: thm RS spec' wenzelm@17336: end wenzelm@17336: fun remove_alls frees thm = Library.foldl (inst_all (sign_of_thm thm)) (thm,frees) wenzelm@17336: fun process_single ((name,atts),rew_imp,frees) args = wenzelm@17336: let wenzelm@17336: fun undo_imps thm = wenzelm@17336: equal_elim (symmetric rew_imp) thm skalberg@14164: haftmann@18358: fun add_final (arg as (thy, thm)) = wenzelm@17336: if name = "" haftmann@18358: then arg |> Library.swap wenzelm@17336: else (writeln (" " ^ name ^ ": " ^ (string_of_thm thm)); haftmann@18358: PureThy.store_thm ((name, thm), []) thy) wenzelm@17336: in wenzelm@17336: args |> apsnd (remove_alls frees) wenzelm@17336: |> apsnd undo_imps wenzelm@17336: |> apsnd standard wenzelm@18729: |> Thm.theory_attributes (map (Attrib.attribute thy) atts) wenzelm@17336: |> add_final haftmann@18358: |> Library.swap wenzelm@17336: end skalberg@14164: wenzelm@17336: fun process_all [proc_arg] args = wenzelm@17336: process_single proc_arg args wenzelm@17336: | process_all (proc_arg::rest) (thy,thm) = wenzelm@17336: let wenzelm@17336: val single_th = thm RS conjunct1 wenzelm@17336: val rest_th = thm RS conjunct2 wenzelm@17336: val (thy',_) = process_single proc_arg (thy,single_th) wenzelm@17336: in wenzelm@17336: process_all rest (thy',rest_th) wenzelm@17336: end wenzelm@17336: | process_all [] _ = error "SpecificationPackage.process_spec internal error" wenzelm@17336: val alt_names = map fst alt_props wenzelm@17336: val _ = if exists (fn(name,_) => not (name = "")) alt_names wenzelm@17336: then writeln "specification" wenzelm@17336: else () wenzelm@17336: in wenzelm@19876: arg |> apsnd Thm.freezeT wenzelm@17336: |> process_all (zip3 alt_names rew_imps frees) wenzelm@17336: end wenzelm@18729: fun post_proc (context, th) = wenzelm@18729: post_process (Context.theory_of context, th) |>> Context.Theory; skalberg@14115: in wenzelm@18799: IsarThy.theorem_i PureThy.internalK wenzelm@18729: ("", [add_spec axiomatic (zip3 names cnames overloaded), post_proc]) wenzelm@19585: (HOLogic.mk_Trueprop ex_prop, []) thy skalberg@14115: end skalberg@14115: wenzelm@17336: skalberg@14115: (* outer syntax *) skalberg@14115: wenzelm@17057: local structure P = OuterParse and K = OuterKeyword in skalberg@14115: skalberg@14116: val opt_name = Scan.optional (P.name --| P.$$$ ":") "" wenzelm@17336: val opt_overloaded = P.opt_keyword "overloaded"; skalberg@14116: skalberg@14115: val specification_decl = wenzelm@17336: P.$$$ "(" |-- Scan.repeat1 (opt_name -- P.term -- opt_overloaded) --| P.$$$ ")" -- wenzelm@17336: Scan.repeat1 (P.opt_thm_name ":" -- P.prop) skalberg@14115: skalberg@14115: val specificationP = skalberg@14115: OuterSyntax.command "specification" "define constants by specification" K.thy_goal skalberg@14164: (specification_decl >> (fn (cos,alt_props) => wenzelm@17336: Toplevel.print o (Toplevel.theory_to_proof wenzelm@17336: (process_spec NONE cos alt_props)))) skalberg@14222: skalberg@14222: val ax_specification_decl = skalberg@14222: P.name -- skalberg@14222: (P.$$$ "(" |-- Scan.repeat1 (opt_name -- P.term -- opt_overloaded) --| P.$$$ ")" -- wenzelm@17336: Scan.repeat1 (P.opt_thm_name ":" -- P.prop)) skalberg@14115: skalberg@14222: val ax_specificationP = skalberg@14222: OuterSyntax.command "ax_specification" "define constants by specification" K.thy_goal skalberg@14222: (ax_specification_decl >> (fn (axname,(cos,alt_props)) => wenzelm@17336: Toplevel.print o (Toplevel.theory_to_proof wenzelm@17336: (process_spec (SOME axname) cos alt_props)))) skalberg@14222: skalberg@14222: val _ = OuterSyntax.add_parsers [specificationP,ax_specificationP] skalberg@14115: skalberg@14115: end skalberg@14115: skalberg@14115: skalberg@14115: end