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