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