src/HOL/Tools/specification_package.ML
author wenzelm
Sun, 07 May 2006 00:22:05 +0200
changeset 19585 70a1ce3b23ae
parent 19414 a21431e996bf
child 19876 11d447d5d68c
permissions -rw-r--r--
removed 'concl is' patterns;
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
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
     5
Package for defining constants by specification.
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
     6
*)
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
     7
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
     8
signature SPECIFICATION_PACKAGE =
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
     9
sig
17336
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
    10
  val quiet_mode: bool ref
18729
216e31270509 simplified type attribute;
wenzelm
parents: 18418
diff changeset
    11
  val add_specification: string option -> (bstring * xstring * bool) list ->
216e31270509 simplified type attribute;
wenzelm
parents: 18418
diff changeset
    12
    theory * thm -> theory * thm
14115
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
    13
end
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
    14
17895
6274b426594b moved helper lemma to HilbertChoice.thy;
wenzelm
parents: 17377
diff changeset
    15
structure SpecificationPackage: SPECIFICATION_PACKAGE =
14115
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
    16
struct
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
    17
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
    18
(* messages *)
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
    19
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
    20
val quiet_mode = ref false
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
    21
fun message s = if ! quiet_mode then () else writeln s
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
    22
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
    23
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
    24
(* Actual code *)
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
    25
14222
1e61f23fd359 Added axiomatic specifications (ax_specification).
skalberg
parents: 14167
diff changeset
    26
local
17895
6274b426594b moved helper lemma to HilbertChoice.thy;
wenzelm
parents: 17377
diff changeset
    27
    val exE_some = thm "exE_some";
6274b426594b moved helper lemma to HilbertChoice.thy;
wenzelm
parents: 17377
diff changeset
    28
14222
1e61f23fd359 Added axiomatic specifications (ax_specification).
skalberg
parents: 14167
diff changeset
    29
    fun mk_definitional [] arg = arg
1e61f23fd359 Added axiomatic specifications (ax_specification).
skalberg
parents: 14167
diff changeset
    30
      | mk_definitional ((thname,cname,covld)::cos) (thy,thm) =
17336
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
    31
        case HOLogic.dest_Trueprop (concl_of thm) of
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
    32
            Const("Ex",_) $ P =>
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
    33
            let
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
    34
                val ctype = domain_type (type_of P)
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
    35
                val cname_full = Sign.intern_const (sign_of thy) cname
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
    36
                val cdefname = if thname = ""
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
    37
                               then Thm.def_name (Sign.base_name cname)
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
    38
                               else thname
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
    39
                val def_eq = Logic.mk_equals (Const(cname_full,ctype),
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
    40
                                              HOLogic.choice_const ctype $  P)
18358
0a733e11021a re-oriented some result tuples in PureThy
haftmann
parents: 17895
diff changeset
    41
                val (thms, thy') = PureThy.add_defs_i covld [((cdefname,def_eq),[])] thy
17895
6274b426594b moved helper lemma to HilbertChoice.thy;
wenzelm
parents: 17377
diff changeset
    42
                val thm' = [thm,hd thms] MRS exE_some
17336
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
    43
            in
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
    44
                mk_definitional cos (thy',thm')
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
    45
            end
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
    46
          | _ => raise THM ("Internal error: Bad specification theorem",0,[thm])
14222
1e61f23fd359 Added axiomatic specifications (ax_specification).
skalberg
parents: 14167
diff changeset
    47
1e61f23fd359 Added axiomatic specifications (ax_specification).
skalberg
parents: 14167
diff changeset
    48
    fun mk_axiomatic axname cos arg =
17336
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
    49
        let
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
    50
            fun process [] (thy,tm) =
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
    51
                let
18377
0e1d025d57b3 oriented result pairs in PureThy
haftmann
parents: 18358
diff changeset
    52
                    val (thms, thy') = PureThy.add_axioms_i [((axname,HOLogic.mk_Trueprop tm),[])] thy
17336
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
    53
                in
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
    54
                    (thy',hd thms)
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
    55
                end
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
    56
              | process ((thname,cname,covld)::cos) (thy,tm) =
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
    57
                case tm of
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
    58
                    Const("Ex",_) $ P =>
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
    59
                    let
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
    60
                        val ctype = domain_type (type_of P)
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
    61
                        val cname_full = Sign.intern_const (sign_of thy) cname
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
    62
                        val cdefname = if thname = ""
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
    63
                                       then Thm.def_name (Sign.base_name cname)
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
    64
                                       else thname
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
    65
                        val co = Const(cname_full,ctype)
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
    66
                        val thy' = Theory.add_finals_i covld [co] thy
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
    67
                        val tm' = case P of
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
    68
                                      Abs(_, _, bodt) => subst_bound (co, bodt)
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
    69
                                    | _ => P $ co
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
    70
                    in
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
    71
                        process cos (thy',tm')
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
    72
                    end
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
    73
                  | _ => raise TERM ("Internal error: Bad specification theorem",[tm])
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
    74
        in
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
    75
            process cos arg
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
    76
        end
14115
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
    77
14222
1e61f23fd359 Added axiomatic specifications (ax_specification).
skalberg
parents: 14167
diff changeset
    78
in
1e61f23fd359 Added axiomatic specifications (ax_specification).
skalberg
parents: 14167
diff changeset
    79
fun proc_exprop axiomatic cos arg =
1e61f23fd359 Added axiomatic specifications (ax_specification).
skalberg
parents: 14167
diff changeset
    80
    case axiomatic of
17336
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
    81
        SOME axname => mk_axiomatic axname cos (apsnd (HOLogic.dest_Trueprop o concl_of) arg)
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14769
diff changeset
    82
      | NONE => mk_definitional cos arg
14222
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
18729
216e31270509 simplified type attribute;
wenzelm
parents: 18418
diff changeset
    85
fun add_specification axiomatic cos arg =
14121
d2a0fd183f5f Added handling of free variables (provided they are of sort HOL.type).
skalberg
parents: 14118
diff changeset
    86
    arg |> apsnd freezeT
17336
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
    87
        |> proc_exprop axiomatic cos
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
    88
        |> apsnd standard
14115
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
    89
18729
216e31270509 simplified type attribute;
wenzelm
parents: 18418
diff changeset
    90
fun add_spec x y (context, thm) =
216e31270509 simplified type attribute;
wenzelm
parents: 18418
diff changeset
    91
  (Context.the_theory context, thm) |> add_specification x y |>> Context.Theory;
216e31270509 simplified type attribute;
wenzelm
parents: 18418
diff changeset
    92
216e31270509 simplified type attribute;
wenzelm
parents: 18418
diff changeset
    93
14115
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
    94
(* Collect all intances of constants in term *)
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
    95
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
    96
fun collect_consts (        t $ u,tms) = collect_consts (u,collect_consts (t,tms))
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
    97
  | collect_consts (   Abs(_,_,t),tms) = collect_consts (t,tms)
18921
f47c46d7d654 canonical member/insert/merge;
wenzelm
parents: 18799
diff changeset
    98
  | collect_consts (tm as Const _,tms) = insert (op aconv) tm tms
14115
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
    99
  | collect_consts (            _,tms) = tms
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
   100
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
   101
(* Complementing Type.varify... *)
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 unvarify t fmap =
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
   104
    let
17336
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   105
        val fmap' = map Library.swap fmap
17377
afaa031ed4da introduced AList.lookup
haftmann
parents: 17336
diff changeset
   106
        fun unthaw (f as (a, S)) =
afaa031ed4da introduced AList.lookup
haftmann
parents: 17336
diff changeset
   107
            (case AList.lookup (op =) fmap' a of
17336
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   108
                 NONE => TVar f
17377
afaa031ed4da introduced AList.lookup
haftmann
parents: 17336
diff changeset
   109
               | SOME (b, _) => TFree (b, S))
14115
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
   110
    in
17336
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   111
        map_term_types (map_type_tvar unthaw) t
14115
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
   112
    end
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
   113
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
   114
(* The syntactic meddling needed to setup add_specification for work *)
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
   115
17336
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   116
fun process_spec axiomatic cos alt_props thy =
14115
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
   117
    let
17336
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   118
        fun zip3 [] [] [] = []
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   119
          | zip3 (x::xs) (y::ys) (z::zs) = (x,y,z)::zip3 xs ys zs
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   120
          | zip3 _ _ _ = error "SpecificationPackage.process_spec internal error"
14167
3c29bba24aa4 Improved the error messages (slightly).
skalberg
parents: 14166
diff changeset
   121
17336
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   122
        fun myfoldr f [x] = x
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   123
          | myfoldr f (x::xs) = f (x,myfoldr f xs)
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   124
          | myfoldr f [] = error "SpecificationPackage.process_spec internal error"
14164
8c3fab596219 Allowed for splitting the specification over several lemmas.
skalberg
parents: 14122
diff changeset
   125
17895
6274b426594b moved helper lemma to HilbertChoice.thy;
wenzelm
parents: 17377
diff changeset
   126
        val rew_imps = alt_props |>
18776
fdc5379fd359 ObjectLogic.atomize_cterm;
wenzelm
parents: 18729
diff changeset
   127
          map (ObjectLogic.atomize_cterm o Thm.read_cterm thy o rpair Term.propT o snd)
17895
6274b426594b moved helper lemma to HilbertChoice.thy;
wenzelm
parents: 17377
diff changeset
   128
        val props' = rew_imps |>
6274b426594b moved helper lemma to HilbertChoice.thy;
wenzelm
parents: 17377
diff changeset
   129
          map (HOLogic.dest_Trueprop o term_of o snd o dest_equals o cprop_of)
14166
1ed8f955727d Cleaned up the code.
skalberg
parents: 14164
diff changeset
   130
17336
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   131
        fun proc_single prop =
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   132
            let
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   133
                val frees = Term.term_frees prop
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   134
                val tsig = Sign.tsig_of (sign_of thy)
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   135
                val _ = assert (forall (fn v => Type.of_sort tsig (type_of v,HOLogic.typeS)) frees)
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   136
                               "Specificaton: Only free variables of sort 'type' allowed"
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   137
                val prop_closed = foldr (fn ((vname,T),prop) => HOLogic.mk_all (vname,T,prop)) prop (map dest_Free frees)
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   138
            in
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   139
                (prop_closed,frees)
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   140
            end
14166
1ed8f955727d Cleaned up the code.
skalberg
parents: 14164
diff changeset
   141
17336
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   142
        val props'' = map proc_single props'
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   143
        val frees = map snd props''
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   144
        val prop  = myfoldr HOLogic.mk_conj (map fst props'')
17895
6274b426594b moved helper lemma to HilbertChoice.thy;
wenzelm
parents: 17377
diff changeset
   145
        val cprop = cterm_of thy (HOLogic.mk_Trueprop prop)
14164
8c3fab596219 Allowed for splitting the specification over several lemmas.
skalberg
parents: 14122
diff changeset
   146
17336
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   147
        val (prop_thawed,vmap) = Type.varify (prop,[])
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   148
        val thawed_prop_consts = collect_consts (prop_thawed,[])
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   149
        val (altcos,overloaded) = Library.split_list cos
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   150
        val (names,sconsts) = Library.split_list altcos
17895
6274b426594b moved helper lemma to HilbertChoice.thy;
wenzelm
parents: 17377
diff changeset
   151
        val consts = map (term_of o Thm.read_cterm thy o rpair TypeInfer.logicT) sconsts
17336
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   152
        val _ = assert (not (Library.exists (not o Term.is_Const) consts))
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   153
                       "Specification: Non-constant found as parameter"
14166
1ed8f955727d Cleaned up the code.
skalberg
parents: 14164
diff changeset
   154
17336
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   155
        fun proc_const c =
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   156
            let
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   157
                val c' = fst (Type.varify (c,[]))
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   158
                val (cname,ctyp) = dest_Const c'
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   159
            in
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   160
                case List.filter (fn t => let val (name,typ) = dest_Const t
19414
a21431e996bf Sign.typ_equiv;
wenzelm
parents: 18921
diff changeset
   161
                                     in name = cname andalso Sign.typ_equiv thy (typ, ctyp)
17336
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   162
                                     end) thawed_prop_consts of
17895
6274b426594b moved helper lemma to HilbertChoice.thy;
wenzelm
parents: 17377
diff changeset
   163
                    [] => error ("Specification: No suitable instances of constant \"" ^ (Sign.string_of_term thy c) ^ "\" found")
17336
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   164
                  | [cf] => unvarify cf vmap
17895
6274b426594b moved helper lemma to HilbertChoice.thy;
wenzelm
parents: 17377
diff changeset
   165
                  | _ => error ("Specification: Several variations of \"" ^ (Sign.string_of_term thy c) ^ "\" found (try applying explicit type constraints)")
17336
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   166
            end
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   167
        val proc_consts = map proc_const consts
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   168
        fun mk_exist (c,prop) =
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   169
            let
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   170
                val T = type_of c
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   171
                val cname = Sign.base_name (fst (dest_Const c))
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   172
                val vname = if Syntax.is_identifier cname
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   173
                            then cname
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   174
                            else "x"
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   175
            in
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   176
                HOLogic.exists_const T $ Abs(vname,T,Term.abstract_over (c,prop))
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   177
            end
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   178
        val ex_prop = foldr mk_exist prop proc_consts
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   179
        val cnames = map (fst o dest_Const) proc_consts
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   180
        fun post_process (arg as (thy,thm)) =
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   181
            let
17895
6274b426594b moved helper lemma to HilbertChoice.thy;
wenzelm
parents: 17377
diff changeset
   182
                fun inst_all thy (thm,v) =
17336
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   183
                    let
17895
6274b426594b moved helper lemma to HilbertChoice.thy;
wenzelm
parents: 17377
diff changeset
   184
                        val cv = cterm_of thy v
17336
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   185
                        val cT = ctyp_of_term cv
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   186
                        val spec' = instantiate' [SOME cT] [NONE,SOME cv] spec
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   187
                    in
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   188
                        thm RS spec'
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   189
                    end
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   190
                fun remove_alls frees thm = Library.foldl (inst_all (sign_of_thm thm)) (thm,frees)
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   191
                fun process_single ((name,atts),rew_imp,frees) args =
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   192
                    let
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   193
                        fun undo_imps thm =
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   194
                            equal_elim (symmetric rew_imp) thm
14164
8c3fab596219 Allowed for splitting the specification over several lemmas.
skalberg
parents: 14122
diff changeset
   195
18358
0a733e11021a re-oriented some result tuples in PureThy
haftmann
parents: 17895
diff changeset
   196
                        fun add_final (arg as (thy, thm)) =
17336
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   197
                            if name = ""
18358
0a733e11021a re-oriented some result tuples in PureThy
haftmann
parents: 17895
diff changeset
   198
                            then arg |> Library.swap
17336
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   199
                            else (writeln ("  " ^ name ^ ": " ^ (string_of_thm thm));
18358
0a733e11021a re-oriented some result tuples in PureThy
haftmann
parents: 17895
diff changeset
   200
                                  PureThy.store_thm ((name, thm), []) thy)
17336
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   201
                    in
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   202
                        args |> apsnd (remove_alls frees)
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   203
                             |> apsnd undo_imps
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   204
                             |> apsnd standard
18729
216e31270509 simplified type attribute;
wenzelm
parents: 18418
diff changeset
   205
                             |> Thm.theory_attributes (map (Attrib.attribute thy) atts)
17336
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   206
                             |> add_final
18358
0a733e11021a re-oriented some result tuples in PureThy
haftmann
parents: 17895
diff changeset
   207
                             |> Library.swap
17336
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   208
                    end
14164
8c3fab596219 Allowed for splitting the specification over several lemmas.
skalberg
parents: 14122
diff changeset
   209
17336
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   210
                fun process_all [proc_arg] args =
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   211
                    process_single proc_arg args
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   212
                  | process_all (proc_arg::rest) (thy,thm) =
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   213
                    let
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   214
                        val single_th = thm RS conjunct1
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   215
                        val rest_th   = thm RS conjunct2
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   216
                        val (thy',_)  = process_single proc_arg (thy,single_th)
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   217
                    in
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   218
                        process_all rest (thy',rest_th)
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   219
                    end
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   220
                  | process_all [] _ = error "SpecificationPackage.process_spec internal error"
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   221
                val alt_names = map fst alt_props
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   222
                val _ = if exists (fn(name,_) => not (name = "")) alt_names
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   223
                        then writeln "specification"
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   224
                        else ()
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   225
            in
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   226
                arg |> apsnd freezeT
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   227
                    |> process_all (zip3 alt_names rew_imps frees)
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   228
            end
18729
216e31270509 simplified type attribute;
wenzelm
parents: 18418
diff changeset
   229
            fun post_proc (context, th) =
216e31270509 simplified type attribute;
wenzelm
parents: 18418
diff changeset
   230
                post_process (Context.theory_of context, th) |>> Context.Theory;
14115
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
   231
    in
18799
f137c5e971f5 moved theorem tags from Drule to PureThy;
wenzelm
parents: 18776
diff changeset
   232
      IsarThy.theorem_i PureThy.internalK
18729
216e31270509 simplified type attribute;
wenzelm
parents: 18418
diff changeset
   233
        ("", [add_spec axiomatic (zip3 names cnames overloaded), post_proc])
19585
70a1ce3b23ae removed 'concl is' patterns;
wenzelm
parents: 19414
diff changeset
   234
        (HOLogic.mk_Trueprop ex_prop, []) thy
14115
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
   235
    end
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
   236
17336
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   237
14115
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
   238
(* outer syntax *)
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
   239
17057
0934ac31985f OuterKeyword;
wenzelm
parents: 15945
diff changeset
   240
local structure P = OuterParse and K = OuterKeyword in
14115
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
   241
14116
63337d8e6e0f Added optional theorem names for the constant definitions added during
skalberg
parents: 14115
diff changeset
   242
val opt_name = Scan.optional (P.name --| P.$$$ ":") ""
17336
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   243
val opt_overloaded = P.opt_keyword "overloaded";
14116
63337d8e6e0f Added optional theorem names for the constant definitions added during
skalberg
parents: 14115
diff changeset
   244
14115
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
   245
val specification_decl =
17336
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   246
  P.$$$ "(" |-- Scan.repeat1 (opt_name -- P.term -- opt_overloaded) --| P.$$$ ")" --
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   247
          Scan.repeat1 (P.opt_thm_name ":" -- P.prop)
14115
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
   248
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
   249
val specificationP =
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
   250
  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
   251
    (specification_decl >> (fn (cos,alt_props) =>
17336
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   252
                               Toplevel.print o (Toplevel.theory_to_proof
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   253
                                                     (process_spec NONE cos alt_props))))
14222
1e61f23fd359 Added axiomatic specifications (ax_specification).
skalberg
parents: 14167
diff changeset
   254
1e61f23fd359 Added axiomatic specifications (ax_specification).
skalberg
parents: 14167
diff changeset
   255
val ax_specification_decl =
1e61f23fd359 Added axiomatic specifications (ax_specification).
skalberg
parents: 14167
diff changeset
   256
    P.name --
1e61f23fd359 Added axiomatic specifications (ax_specification).
skalberg
parents: 14167
diff changeset
   257
    (P.$$$ "(" |-- Scan.repeat1 (opt_name -- P.term -- opt_overloaded) --| P.$$$ ")" --
17336
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   258
           Scan.repeat1 (P.opt_thm_name ":" -- P.prop))
14115
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
   259
14222
1e61f23fd359 Added axiomatic specifications (ax_specification).
skalberg
parents: 14167
diff changeset
   260
val ax_specificationP =
1e61f23fd359 Added axiomatic specifications (ax_specification).
skalberg
parents: 14167
diff changeset
   261
  OuterSyntax.command "ax_specification" "define constants by specification" K.thy_goal
1e61f23fd359 Added axiomatic specifications (ax_specification).
skalberg
parents: 14167
diff changeset
   262
    (ax_specification_decl >> (fn (axname,(cos,alt_props)) =>
17336
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   263
                               Toplevel.print o (Toplevel.theory_to_proof
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   264
                                                     (process_spec (SOME axname) cos alt_props))))
14222
1e61f23fd359 Added axiomatic specifications (ax_specification).
skalberg
parents: 14167
diff changeset
   265
1e61f23fd359 Added axiomatic specifications (ax_specification).
skalberg
parents: 14167
diff changeset
   266
val _ = OuterSyntax.add_parsers [specificationP,ax_specificationP]
14115
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
   267
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
   268
end
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
   269
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
   270
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
   271
end