src/HOL/Tools/specification_package.ML
author wenzelm
Sat, 06 Oct 2007 16:50:04 +0200
changeset 24867 e5b55d7be9bb
parent 24707 dfeb98f84e93
child 26478 9d1029ce0e13
permissions -rw-r--r--
simplified interfaces for outer syntax;
     1 (*  Title:      HOL/Tools/specification_package.ML
     2     ID:         $Id$
     3     Author:     Sebastian Skalberg, TU Muenchen
     4 
     5 Package for defining constants by specification.
     6 *)
     7 
     8 signature SPECIFICATION_PACKAGE =
     9 sig
    10   val quiet_mode: bool ref
    11   val add_specification: string option -> (bstring * xstring * bool) list ->
    12     theory * thm -> theory * thm
    13 end
    14 
    15 structure SpecificationPackage: SPECIFICATION_PACKAGE =
    16 struct
    17 
    18 (* messages *)
    19 
    20 val quiet_mode = ref false
    21 fun message s = if ! quiet_mode then () else writeln s
    22 
    23 
    24 (* Actual code *)
    25 
    26 local
    27     val exE_some = thm "exE_some";
    28 
    29     fun mk_definitional [] arg = arg
    30       | mk_definitional ((thname,cname,covld)::cos) (thy,thm) =
    31         case HOLogic.dest_Trueprop (concl_of thm) of
    32             Const("Ex",_) $ P =>
    33             let
    34                 val ctype = domain_type (type_of P)
    35                 val cname_full = Sign.intern_const thy cname
    36                 val cdefname = if thname = ""
    37                                then Thm.def_name (Sign.base_name cname)
    38                                else thname
    39                 val def_eq = Logic.mk_equals (Const(cname_full,ctype),
    40                                               HOLogic.choice_const ctype $  P)
    41                 val (thms, thy') = PureThy.add_defs_i covld [((cdefname,def_eq),[])] thy
    42                 val thm' = [thm,hd thms] MRS exE_some
    43             in
    44                 mk_definitional cos (thy',thm')
    45             end
    46           | _ => raise THM ("Internal error: Bad specification theorem",0,[thm])
    47 
    48     fun mk_axiomatic axname cos arg =
    49         let
    50             fun process [] (thy,tm) =
    51                 let
    52                     val (thms, thy') = PureThy.add_axioms_i [((axname,HOLogic.mk_Trueprop tm),[])] thy
    53                 in
    54                     (thy',hd thms)
    55                 end
    56               | process ((thname,cname,covld)::cos) (thy,tm) =
    57                 case tm of
    58                     Const("Ex",_) $ P =>
    59                     let
    60                         val ctype = domain_type (type_of P)
    61                         val cname_full = Sign.intern_const thy cname
    62                         val cdefname = if thname = ""
    63                                        then Thm.def_name (Sign.base_name cname)
    64                                        else thname
    65                         val co = Const(cname_full,ctype)
    66                         val thy' = Theory.add_finals_i covld [co] thy
    67                         val tm' = case P of
    68                                       Abs(_, _, bodt) => subst_bound (co, bodt)
    69                                     | _ => P $ co
    70                     in
    71                         process cos (thy',tm')
    72                     end
    73                   | _ => raise TERM ("Internal error: Bad specification theorem",[tm])
    74         in
    75             process cos arg
    76         end
    77 
    78 in
    79 fun proc_exprop axiomatic cos arg =
    80     case axiomatic of
    81         SOME axname => mk_axiomatic axname cos (apsnd (HOLogic.dest_Trueprop o concl_of) arg)
    82       | NONE => mk_definitional cos arg
    83 end
    84 
    85 fun add_specification axiomatic cos arg =
    86     arg |> apsnd Thm.freezeT
    87         |> proc_exprop axiomatic cos
    88         |> apsnd standard
    89 
    90 
    91 (* Collect all intances of constants in term *)
    92 
    93 fun collect_consts (        t $ u,tms) = collect_consts (u,collect_consts (t,tms))
    94   | collect_consts (   Abs(_,_,t),tms) = collect_consts (t,tms)
    95   | collect_consts (tm as Const _,tms) = insert (op aconv) tm tms
    96   | collect_consts (            _,tms) = tms
    97 
    98 (* Complementing Type.varify... *)
    99 
   100 fun unvarify t fmap =
   101     let
   102         val fmap' = map Library.swap fmap
   103         fun unthaw (f as (a, S)) =
   104             (case AList.lookup (op =) fmap' a of
   105                  NONE => TVar f
   106                | SOME (b, _) => TFree (b, S))
   107     in
   108         map_types (map_type_tvar unthaw) t
   109     end
   110 
   111 (* The syntactic meddling needed to setup add_specification for work *)
   112 
   113 fun process_spec axiomatic cos alt_props thy =
   114     let
   115         fun zip3 [] [] [] = []
   116           | zip3 (x::xs) (y::ys) (z::zs) = (x,y,z)::zip3 xs ys zs
   117           | zip3 _ _ _ = error "SpecificationPackage.process_spec internal error"
   118 
   119         fun myfoldr f [x] = x
   120           | myfoldr f (x::xs) = f (x,myfoldr f xs)
   121           | myfoldr f [] = error "SpecificationPackage.process_spec internal error"
   122 
   123         val rew_imps = alt_props |>
   124           map (ObjectLogic.atomize o Thm.cterm_of thy o Syntax.read_prop_global thy o snd)
   125         val props' = rew_imps |>
   126           map (HOLogic.dest_Trueprop o term_of o snd o Thm.dest_equals o cprop_of)
   127 
   128         fun proc_single prop =
   129             let
   130                 val frees = Term.term_frees prop
   131                 val _ = forall (fn v => Sign.of_sort thy (type_of v,HOLogic.typeS)) frees
   132                   orelse error "Specificaton: Only free variables of sort 'type' allowed"
   133                 val prop_closed = foldr (fn ((vname,T),prop) => HOLogic.mk_all (vname,T,prop)) prop (map dest_Free frees)
   134             in
   135                 (prop_closed,frees)
   136             end
   137 
   138         val props'' = map proc_single props'
   139         val frees = map snd props''
   140         val prop  = myfoldr HOLogic.mk_conj (map fst props'')
   141         val cprop = cterm_of thy (HOLogic.mk_Trueprop prop)
   142 
   143         val (vmap, prop_thawed) = Type.varify [] prop
   144         val thawed_prop_consts = collect_consts (prop_thawed,[])
   145         val (altcos,overloaded) = Library.split_list cos
   146         val (names,sconsts) = Library.split_list altcos
   147         val consts = map (Syntax.read_term_global thy) sconsts
   148         val _ = not (Library.exists (not o Term.is_Const) consts)
   149           orelse error "Specification: Non-constant found as parameter"
   150 
   151         fun proc_const c =
   152             let
   153                 val (_, c') = Type.varify [] c
   154                 val (cname,ctyp) = dest_Const c'
   155             in
   156                 case List.filter (fn t => let val (name,typ) = dest_Const t
   157                                      in name = cname andalso Sign.typ_equiv thy (typ, ctyp)
   158                                      end) thawed_prop_consts of
   159                     [] => error ("Specification: No suitable instances of constant \"" ^ (Sign.string_of_term thy c) ^ "\" found")
   160                   | [cf] => unvarify cf vmap
   161                   | _ => error ("Specification: Several variations of \"" ^ (Sign.string_of_term thy c) ^ "\" found (try applying explicit type constraints)")
   162             end
   163         val proc_consts = map proc_const consts
   164         fun mk_exist (c,prop) =
   165             let
   166                 val T = type_of c
   167                 val cname = Sign.base_name (fst (dest_Const c))
   168                 val vname = if Syntax.is_identifier cname
   169                             then cname
   170                             else "x"
   171             in
   172                 HOLogic.exists_const T $ Abs(vname,T,Term.abstract_over (c,prop))
   173             end
   174         val ex_prop = foldr mk_exist prop proc_consts
   175         val cnames = map (fst o dest_Const) proc_consts
   176         fun post_process (arg as (thy,thm)) =
   177             let
   178                 fun inst_all thy (thm,v) =
   179                     let
   180                         val cv = cterm_of thy v
   181                         val cT = ctyp_of_term cv
   182                         val spec' = instantiate' [SOME cT] [NONE,SOME cv] spec
   183                     in
   184                         thm RS spec'
   185                     end
   186                 fun remove_alls frees thm =
   187                     Library.foldl (inst_all (Thm.theory_of_thm thm)) (thm,frees)
   188                 fun process_single ((name,atts),rew_imp,frees) args =
   189                     let
   190                         fun undo_imps thm =
   191                             equal_elim (symmetric rew_imp) thm
   192 
   193                         fun add_final (arg as (thy, thm)) =
   194                             if name = ""
   195                             then arg |> Library.swap
   196                             else (writeln ("  " ^ name ^ ": " ^ (string_of_thm thm));
   197                                   PureThy.store_thm ((name, thm), []) thy)
   198                     in
   199                         args |> apsnd (remove_alls frees)
   200                              |> apsnd undo_imps
   201                              |> apsnd standard
   202                              |> Thm.theory_attributes (map (Attrib.attribute thy) atts)
   203                              |> add_final
   204                              |> Library.swap
   205                     end
   206 
   207                 fun process_all [proc_arg] args =
   208                     process_single proc_arg args
   209                   | process_all (proc_arg::rest) (thy,thm) =
   210                     let
   211                         val single_th = thm RS conjunct1
   212                         val rest_th   = thm RS conjunct2
   213                         val (thy',_)  = process_single proc_arg (thy,single_th)
   214                     in
   215                         process_all rest (thy',rest_th)
   216                     end
   217                   | process_all [] _ = error "SpecificationPackage.process_spec internal error"
   218                 val alt_names = map fst alt_props
   219                 val _ = if exists (fn(name,_) => not (name = "")) alt_names
   220                         then writeln "specification"
   221                         else ()
   222             in
   223                 arg |> apsnd Thm.freezeT
   224                     |> process_all (zip3 alt_names rew_imps frees)
   225             end
   226 
   227       fun after_qed [[thm]] = ProofContext.theory (fn thy =>
   228         #1 (post_process (add_specification axiomatic (zip3 names cnames overloaded) (thy, thm))));
   229     in
   230       thy
   231       |> ProofContext.init
   232       |> Proof.theorem_i NONE after_qed [[(HOLogic.mk_Trueprop ex_prop, [])]]
   233     end;
   234 
   235 
   236 (* outer syntax *)
   237 
   238 local structure P = OuterParse and K = OuterKeyword in
   239 
   240 val opt_name = Scan.optional (P.name --| P.$$$ ":") ""
   241 val opt_overloaded = P.opt_keyword "overloaded";
   242 
   243 val specification_decl =
   244   P.$$$ "(" |-- Scan.repeat1 (opt_name -- P.term -- opt_overloaded) --| P.$$$ ")" --
   245           Scan.repeat1 (SpecParse.opt_thm_name ":" -- P.prop)
   246 
   247 val _ =
   248   OuterSyntax.command "specification" "define constants by specification" K.thy_goal
   249     (specification_decl >> (fn (cos,alt_props) =>
   250                                Toplevel.print o (Toplevel.theory_to_proof
   251                                                      (process_spec NONE cos alt_props))))
   252 
   253 val ax_specification_decl =
   254     P.name --
   255     (P.$$$ "(" |-- Scan.repeat1 (opt_name -- P.term -- opt_overloaded) --| P.$$$ ")" --
   256            Scan.repeat1 (SpecParse.opt_thm_name ":" -- P.prop))
   257 
   258 val _ =
   259   OuterSyntax.command "ax_specification" "define constants by specification" K.thy_goal
   260     (ax_specification_decl >> (fn (axname,(cos,alt_props)) =>
   261                                Toplevel.print o (Toplevel.theory_to_proof
   262                                                      (process_spec (SOME axname) cos alt_props))))
   263 
   264 end
   265 
   266 
   267 end