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