src/HOL/Tools/specification_package.ML
author skalberg
Thu, 09 Oct 2003 18:13:32 +0200
changeset 14223 0ee05eef881b
parent 14222 1e61f23fd359
child 14620 1be590fd2422
permissions -rw-r--r--
Added support for making constants final, that is, ensuring that no definition can be given later (useful for constants whose behaviour is fixed axiomatically rather than definitionally).

(*  Title:      HOL/Tools/specification_package.ML
    ID:         $Id$
    Author:     Sebastian Skalberg, TU Muenchen
    License:    GPL (GNU GENERAL PUBLIC LICENSE)

Package for defining constants by specification.
*)

signature SPECIFICATION_PACKAGE =
sig
    val quiet_mode: bool ref
    val add_specification_i: string option -> (bstring * xstring * bool) list
			     -> theory * thm -> theory * thm
end

structure SpecificationPackage : SPECIFICATION_PACKAGE =
struct

(* messages *)

val quiet_mode = ref false
fun message s = if ! quiet_mode then () else writeln s

local
    val _ = Goal "[| Ex P ; c == Eps P |] ==> P c"
    val _ = by (Asm_simp_tac 1)
    val _ = by (rtac (thm "someI_ex") 1)
    val _ = ba 1
in
val helper = Goals.result()
end

(* Akin to HOLogic.exists_const *)
fun choice_const T = Const("Hilbert_Choice.Eps",(T-->HOLogic.boolT)-->T)

(* Actual code *)

local
    fun mk_definitional [] arg = arg
      | mk_definitional ((thname,cname,covld)::cos) (thy,thm) =
	case HOLogic.dest_Trueprop (concl_of thm) of
	    Const("Ex",_) $ P =>
	    let
		val ctype = domain_type (type_of P)
		val cname_full = Sign.intern_const (sign_of thy) cname
		val cdefname = if thname = ""
			       then Thm.def_name (Sign.base_name cname)
			       else thname
		val def_eq = Logic.mk_equals (Const(cname_full,ctype),choice_const ctype $  P)
		val (thy',thms) = PureThy.add_defs_i covld [((cdefname,def_eq),[])] thy
		val thm' = [thm,hd thms] MRS helper
	    in
		mk_definitional cos (thy',thm')
	    end
	  | _ => raise THM ("Internal error: Bad specification theorem",0,[thm])

    fun mk_axiomatic axname cos arg =
	let
	    fun process [] (thy,tm) =
		let
		    val (thy',thms) = PureThy.add_axioms_i [((axname,HOLogic.mk_Trueprop tm),[])] thy
		in
		    (thy',hd thms)
		end
	      | process ((thname,cname,covld)::cos) (thy,tm) =
		case tm of
		    Const("Ex",_) $ P =>
		    let
			val ctype = domain_type (type_of P)
			val cname_full = Sign.intern_const (sign_of thy) cname
			val cdefname = if thname = ""
				       then Thm.def_name (Sign.base_name cname)
				       else thname
			val co = Const(cname_full,ctype)
			val thy' = Theory.add_finals_i covld [co] thy
			val tm' = case P of
				      Abs(_, _, bodt) => subst_bound (co, bodt)
				    | _ => P $ co
		    in
			process cos (thy',tm')
		    end
		  | _ => raise TERM ("Internal error: Bad specification theorem",[tm])
	in
	    process cos arg
	end

in
fun proc_exprop axiomatic cos arg =
    case axiomatic of
	Some axname => mk_axiomatic axname cos (apsnd (HOLogic.dest_Trueprop o concl_of) arg)
      | None => mk_definitional cos arg
end

fun add_specification_i axiomatic cos arg =
    arg |> apsnd freezeT
	|> proc_exprop axiomatic cos
	|> apsnd standard

(* Collect all intances of constants in term *)

fun collect_consts (        t $ u,tms) = collect_consts (u,collect_consts (t,tms))
  | collect_consts (   Abs(_,_,t),tms) = collect_consts (t,tms)
  | collect_consts (tm as Const _,tms) = gen_ins (op aconv) (tm,tms)
  | collect_consts (            _,tms) = tms

(* Complementing Type.varify... *)

fun unvarify t fmap =
    let
	val fmap' = map Library.swap fmap
	fun unthaw (f as (a,S)) =
	    (case assoc (fmap',a) of
		 None => TVar f
	       | Some b => TFree (b,S))
    in
	map_term_types (map_type_tvar unthaw) t
    end

(* The syntactic meddling needed to setup add_specification for work *)

fun process_spec axiomatic cos alt_props int thy =
    let
	fun zip3 [] [] [] = []
	  | zip3 (x::xs) (y::ys) (z::zs) = (x,y,z)::zip3 xs ys zs
	  | zip3 _ _ _ = error "SpecificationPackage.process_spec internal error"

	fun myfoldr f [x] = x
	  | myfoldr f (x::xs) = f (x,myfoldr f xs)
	  | myfoldr f [] = error "SpecificationPackage.process_spec internal error"

	val sg = sign_of thy
	fun typ_equiv t u =
	    let
		val tsig = Sign.tsig_of sg
	    in
		Type.typ_instance(tsig,t,u) andalso
		Type.typ_instance(tsig,u,t)
	    end

	val cprops = map (Thm.read_cterm sg o rpair Term.propT o snd) alt_props
	val ss = empty_ss setmksimps single addsimps [thm "HOL.atomize_imp",thm "HOL.atomize_all"]
	val rew_imps = map (Simplifier.full_rewrite ss) cprops
	val props' = map (HOLogic.dest_Trueprop o term_of o snd o dest_equals o cprop_of) rew_imps

	fun proc_single prop =
	    let
		val frees = Term.term_frees prop
		val tsig = Sign.tsig_of (sign_of thy)
		val _ = assert (forall (fn v => Type.of_sort tsig (type_of v,HOLogic.typeS)) frees)
			       "Specificaton: Only free variables of sort 'type' allowed"
		val prop_closed = foldr (fn ((vname,T),prop) => HOLogic.mk_all (vname,T,prop)) (map dest_Free frees,prop)
	    in
		(prop_closed,frees)
	    end

	val props'' = map proc_single props'
	val frees = map snd props''
	val prop  = myfoldr HOLogic.mk_conj (map fst props'')
	val cprop = cterm_of sg (HOLogic.mk_Trueprop prop)

	val (prop_thawed,vmap) = Type.varify (prop,[])
	val thawed_prop_consts = collect_consts (prop_thawed,[])
	val (altcos,overloaded) = Library.split_list cos
	val (names,sconsts) = Library.split_list altcos
	val consts = map (term_of o Thm.read_cterm sg o rpair TypeInfer.logicT) sconsts
	val _ = assert (not (Library.exists (not o Term.is_Const) consts))
		       "Specification: Non-constant found as parameter"

	fun proc_const c =
	    let
		val c' = fst (Type.varify (c,[]))
		val (cname,ctyp) = dest_Const c'
	    in
		case filter (fn t => let val (name,typ) = dest_Const t
				     in name = cname andalso typ_equiv typ ctyp
				     end) thawed_prop_consts of
		    [] => error ("Specification: No suitable instances of constant \"" ^ (Sign.string_of_term sg c) ^ "\" found")
		  | [cf] => unvarify cf vmap
		  | _ => error ("Specification: Several variations of \"" ^ (Sign.string_of_term sg c) ^ "\" found (try applying explicit type constraints)")
	    end
	val proc_consts = map proc_const consts
	fun mk_exist (c,prop) =
	    let
		val T = type_of c
		val cname = Sign.base_name (fst (dest_Const c))
		val vname = if Syntax.is_identifier cname
			    then cname
			    else "x"
	    in
		HOLogic.exists_const T $ Abs(vname,T,Term.abstract_over (c,prop))
	    end
	val ex_prop = foldr mk_exist (proc_consts,prop)
	val cnames = map (fst o dest_Const) proc_consts
	fun post_process (arg as (thy,thm)) =
	    let
		fun inst_all sg (thm,v) =
		    let
			val cv = cterm_of sg v
			val cT = ctyp_of_term cv
			val spec' = instantiate' [Some cT] [None,Some cv] spec
		    in
			thm RS spec'
		    end
		fun remove_alls frees thm = foldl (inst_all (sign_of_thm thm)) (thm,frees)
		fun process_single ((name,atts),rew_imp,frees) args =
		    let
			fun undo_imps thm =
			    equal_elim (symmetric rew_imp) thm

			fun apply_atts arg = Thm.apply_attributes (arg,map (Attrib.global_attribute thy) atts)
			fun add_final (arg as (thy,thm)) =
			    if name = ""
			    then arg
			    else (writeln ("  " ^ name ^ ": " ^ (string_of_thm thm));
				  PureThy.store_thm((name,thm),[]) thy)
		    in
			args |> apsnd (remove_alls frees)
			     |> apsnd undo_imps
			     |> apsnd standard
			     |> apply_atts
			     |> add_final
		    end

		fun process_all [proc_arg] args =
		    process_single proc_arg args
		  | process_all (proc_arg::rest) (thy,thm) = 
		    let
			val single_th = thm RS conjunct1
			val rest_th   = thm RS conjunct2
			val (thy',_)  = process_single proc_arg (thy,single_th)
		    in
			process_all rest (thy',rest_th)
		    end
		  | process_all [] _ = error "SpecificationPackage.process_spec internal error"
		val alt_names = map fst alt_props
		val _ = if exists (fn(name,_) => not (name = "")) alt_names
			then writeln "specification"
			else ()
	    in
		arg |> apsnd freezeT
		    |> process_all (zip3 alt_names rew_imps frees)
	    end
    in
	IsarThy.theorem_i Drule.internalK
	    (("",[add_specification_i axiomatic (zip3 names cnames overloaded),post_process]),
	     (HOLogic.mk_Trueprop ex_prop,([],[]))) int thy
    end

(* outer syntax *)

local structure P = OuterParse and K = OuterSyntax.Keyword in

(* taken from ~~/Pure/Isar/isar_syn.ML *)
val opt_overloaded =
  Scan.optional (P.$$$ "(" |-- P.!!! ((P.$$$ "overloaded" >> K true) --| P.$$$ ")")) false

val opt_name = Scan.optional (P.name --| P.$$$ ":") ""

val specification_decl =
    P.$$$ "(" |-- Scan.repeat1 (opt_name -- P.term -- opt_overloaded) --| P.$$$ ")" --
	  Scan.repeat1 (P.opt_thm_name ":" -- P.prop)

val specificationP =
  OuterSyntax.command "specification" "define constants by specification" K.thy_goal
    (specification_decl >> (fn (cos,alt_props) =>
			       Toplevel.print o (Toplevel.theory_to_proof
						     (process_spec None cos alt_props))))

val ax_specification_decl =
    P.name --
    (P.$$$ "(" |-- Scan.repeat1 (opt_name -- P.term -- opt_overloaded) --| P.$$$ ")" --
	   Scan.repeat1 (P.opt_thm_name ":" -- P.prop))

val ax_specificationP =
  OuterSyntax.command "ax_specification" "define constants by specification" K.thy_goal
    (ax_specification_decl >> (fn (axname,(cos,alt_props)) =>
			       Toplevel.print o (Toplevel.theory_to_proof
						     (process_spec (Some axname) cos alt_props))))

val _ = OuterSyntax.add_parsers [specificationP,ax_specificationP]

end


end