src/HOL/Tools/choice_specification.ML
author boehmes
Mon, 08 Nov 2010 12:13:44 +0100
changeset 40424 7550b2cba1cb
parent 39557 fe5722fce758
child 42290 b1f544c84040
permissions -rw-r--r--
better modularization: moved SMT configuration options and diagnostics as well as SMT failure and exception into separate structures (both of which are loaded first and consequently are available to other SMT structures)
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
31723
f5cafe803b55 discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents: 30364
diff changeset
     1
(*  Title:      HOL/Tools/choice_specification.ML
14115
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
     2
    Author:     Sebastian Skalberg, TU Muenchen
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
     3
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
     4
Package for defining constants by specification.
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
     5
*)
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
     6
31723
f5cafe803b55 discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents: 30364
diff changeset
     7
signature CHOICE_SPECIFICATION =
14115
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
     8
sig
35807
e4d1b5cbd429 added support for "specification" and "ax_specification" constructs to Nitpick
blanchet
parents: 35625
diff changeset
     9
  val close_form : term -> term
18729
216e31270509 simplified type attribute;
wenzelm
parents: 18418
diff changeset
    10
  val add_specification: string option -> (bstring * xstring * bool) list ->
216e31270509 simplified type attribute;
wenzelm
parents: 18418
diff changeset
    11
    theory * thm -> theory * thm
14115
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
    12
end
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
    13
31723
f5cafe803b55 discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents: 30364
diff changeset
    14
structure Choice_Specification: CHOICE_SPECIFICATION =
14115
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
    15
struct
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
    16
26478
9d1029ce0e13 eliminated quiete_mode ref (not really needed);
wenzelm
parents: 24867
diff changeset
    17
(* actual code *)
14115
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
    18
35807
e4d1b5cbd429 added support for "specification" and "ax_specification" constructs to Nitpick
blanchet
parents: 35625
diff changeset
    19
fun close_form t =
e4d1b5cbd429 added support for "specification" and "ax_specification" constructs to Nitpick
blanchet
parents: 35625
diff changeset
    20
    fold_rev (fn (s, T) => fn t => HOLogic.mk_all (s, T, t))
e4d1b5cbd429 added support for "specification" and "ax_specification" constructs to Nitpick
blanchet
parents: 35625
diff changeset
    21
             (map dest_Free (OldTerm.term_frees t)) t
e4d1b5cbd429 added support for "specification" and "ax_specification" constructs to Nitpick
blanchet
parents: 35625
diff changeset
    22
14222
1e61f23fd359 Added axiomatic specifications (ax_specification).
skalberg
parents: 14167
diff changeset
    23
local
1e61f23fd359 Added axiomatic specifications (ax_specification).
skalberg
parents: 14167
diff changeset
    24
    fun mk_definitional [] arg = arg
1e61f23fd359 Added axiomatic specifications (ax_specification).
skalberg
parents: 14167
diff changeset
    25
      | mk_definitional ((thname,cname,covld)::cos) (thy,thm) =
17336
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
    26
        case HOLogic.dest_Trueprop (concl_of thm) of
38558
32ad17fe2b9c tuned quotes
haftmann
parents: 38549
diff changeset
    27
            Const(@{const_name Ex},_) $ P =>
17336
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
    28
            let
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
    29
                val ctype = domain_type (type_of P)
22578
b0eb5652f210 removed obsolete sign_of/sign_of_thm;
wenzelm
parents: 22101
diff changeset
    30
                val cname_full = Sign.intern_const thy cname
17336
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
    31
                val cdefname = if thname = ""
30364
577edc39b501 moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents: 30280
diff changeset
    32
                               then Thm.def_name (Long_Name.base_name cname)
17336
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
    33
                               else thname
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
    34
                val def_eq = Logic.mk_equals (Const(cname_full,ctype),
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
    35
                                              HOLogic.choice_const ctype $  P)
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 38756
diff changeset
    36
                val (thms, thy') = Global_Theory.add_defs covld [((Binding.name cdefname, def_eq),[])] thy
26478
9d1029ce0e13 eliminated quiete_mode ref (not really needed);
wenzelm
parents: 24867
diff changeset
    37
                val thm' = [thm,hd thms] MRS @{thm exE_some}
17336
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
    38
            in
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
    39
                mk_definitional cos (thy',thm')
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
    40
            end
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
    41
          | _ => raise THM ("Internal error: Bad specification theorem",0,[thm])
14222
1e61f23fd359 Added axiomatic specifications (ax_specification).
skalberg
parents: 14167
diff changeset
    42
1e61f23fd359 Added axiomatic specifications (ax_specification).
skalberg
parents: 14167
diff changeset
    43
    fun mk_axiomatic axname cos arg =
17336
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
    44
        let
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
    45
            fun process [] (thy,tm) =
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
    46
                let
35897
8758895ea413 eliminated old-style Drule.add_axiom in favour of Specification.axiom, with explicit Drule.export_without_context to imitate the old behaviour;
wenzelm
parents: 35856
diff changeset
    47
                    val (thm, thy') =
8758895ea413 eliminated old-style Drule.add_axiom in favour of Specification.axiom, with explicit Drule.export_without_context to imitate the old behaviour;
wenzelm
parents: 35856
diff changeset
    48
                      Specification.axiom ((Binding.name axname, []), HOLogic.mk_Trueprop tm) thy
17336
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
    49
                in
35897
8758895ea413 eliminated old-style Drule.add_axiom in favour of Specification.axiom, with explicit Drule.export_without_context to imitate the old behaviour;
wenzelm
parents: 35856
diff changeset
    50
                    (thy', Drule.export_without_context thm)
17336
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
    51
                end
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
    52
              | process ((thname,cname,covld)::cos) (thy,tm) =
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
    53
                case tm of
38558
32ad17fe2b9c tuned quotes
haftmann
parents: 38549
diff changeset
    54
                    Const(@{const_name Ex},_) $ P =>
17336
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
    55
                    let
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
    56
                        val ctype = domain_type (type_of P)
22578
b0eb5652f210 removed obsolete sign_of/sign_of_thm;
wenzelm
parents: 22101
diff changeset
    57
                        val cname_full = Sign.intern_const thy cname
17336
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
    58
                        val cdefname = if thname = ""
30364
577edc39b501 moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents: 30280
diff changeset
    59
                                       then Thm.def_name (Long_Name.base_name cname)
17336
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
    60
                                       else thname
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
    61
                        val co = Const(cname_full,ctype)
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
    62
                        val thy' = Theory.add_finals_i covld [co] thy
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
    63
                        val tm' = case P of
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
    64
                                      Abs(_, _, bodt) => subst_bound (co, bodt)
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
    65
                                    | _ => P $ co
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
    66
                    in
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
    67
                        process cos (thy',tm')
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
    68
                    end
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
    69
                  | _ => raise TERM ("Internal error: Bad specification theorem",[tm])
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 arg
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
    72
        end
14115
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
    73
14222
1e61f23fd359 Added axiomatic specifications (ax_specification).
skalberg
parents: 14167
diff changeset
    74
in
1e61f23fd359 Added axiomatic specifications (ax_specification).
skalberg
parents: 14167
diff changeset
    75
fun proc_exprop axiomatic cos arg =
1e61f23fd359 Added axiomatic specifications (ax_specification).
skalberg
parents: 14167
diff changeset
    76
    case axiomatic of
17336
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
    77
        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
    78
      | NONE => mk_definitional cos arg
14222
1e61f23fd359 Added axiomatic specifications (ax_specification).
skalberg
parents: 14167
diff changeset
    79
end
1e61f23fd359 Added axiomatic specifications (ax_specification).
skalberg
parents: 14167
diff changeset
    80
36618
7a0990473e03 specification goal: eliminated redundant Thm.legacy_freezeT -- the goal is properly declared and should always produce fixed types in the result;
wenzelm
parents: 36615
diff changeset
    81
fun add_specification axiomatic cos =
7a0990473e03 specification goal: eliminated redundant Thm.legacy_freezeT -- the goal is properly declared and should always produce fixed types in the result;
wenzelm
parents: 36615
diff changeset
    82
    proc_exprop axiomatic cos
7a0990473e03 specification goal: eliminated redundant Thm.legacy_freezeT -- the goal is properly declared and should always produce fixed types in the result;
wenzelm
parents: 36615
diff changeset
    83
    #> apsnd Drule.export_without_context
14115
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
    84
18729
216e31270509 simplified type attribute;
wenzelm
parents: 18418
diff changeset
    85
14115
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
    86
(* Collect all intances of constants in term *)
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
    87
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
    88
fun collect_consts (        t $ u,tms) = collect_consts (u,collect_consts (t,tms))
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
    89
  | collect_consts (   Abs(_,_,t),tms) = collect_consts (t,tms)
18921
f47c46d7d654 canonical member/insert/merge;
wenzelm
parents: 18799
diff changeset
    90
  | collect_consts (tm as Const _,tms) = insert (op aconv) tm tms
14115
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
    91
  | collect_consts (            _,tms) = tms
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
    92
35845
e5980f0ad025 renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
wenzelm
parents: 35807
diff changeset
    93
(* Complementing Type.varify_global... *)
14115
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
    94
35845
e5980f0ad025 renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
wenzelm
parents: 35807
diff changeset
    95
fun unvarify_global t fmap =
14115
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
    96
    let
17336
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
    97
        val fmap' = map Library.swap fmap
17377
afaa031ed4da introduced AList.lookup
haftmann
parents: 17336
diff changeset
    98
        fun unthaw (f as (a, S)) =
afaa031ed4da introduced AList.lookup
haftmann
parents: 17336
diff changeset
    99
            (case AList.lookup (op =) fmap' a of
17336
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   100
                 NONE => TVar f
17377
afaa031ed4da introduced AList.lookup
haftmann
parents: 17336
diff changeset
   101
               | SOME (b, _) => TFree (b, S))
14115
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
   102
    in
20548
8ef25fe585a8 renamed Term.map_term_types to Term.map_types (cf. Term.fold_types);
wenzelm
parents: 20344
diff changeset
   103
        map_types (map_type_tvar unthaw) t
14115
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
   104
    end
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
   105
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
   106
(* The syntactic meddling needed to setup add_specification for work *)
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
   107
17336
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   108
fun process_spec axiomatic cos alt_props thy =
14115
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
   109
    let
17336
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   110
        fun zip3 [] [] [] = []
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   111
          | zip3 (x::xs) (y::ys) (z::zs) = (x,y,z)::zip3 xs ys zs
31723
f5cafe803b55 discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents: 30364
diff changeset
   112
          | zip3 _ _ _ = error "Choice_Specification.process_spec internal error"
14167
3c29bba24aa4 Improved the error messages (slightly).
skalberg
parents: 14166
diff changeset
   113
17336
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   114
        fun myfoldr f [x] = x
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   115
          | myfoldr f (x::xs) = f (x,myfoldr f xs)
31723
f5cafe803b55 discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents: 30364
diff changeset
   116
          | myfoldr f [] = error "Choice_Specification.process_spec internal error"
14164
8c3fab596219 Allowed for splitting the specification over several lemmas.
skalberg
parents: 14122
diff changeset
   117
17895
6274b426594b moved helper lemma to HilbertChoice.thy;
wenzelm
parents: 17377
diff changeset
   118
        val rew_imps = alt_props |>
35625
9c818cab0dd0 modernized structure Object_Logic;
wenzelm
parents: 35021
diff changeset
   119
          map (Object_Logic.atomize o Thm.cterm_of thy o Syntax.read_prop_global thy o snd)
17895
6274b426594b moved helper lemma to HilbertChoice.thy;
wenzelm
parents: 17377
diff changeset
   120
        val props' = rew_imps |>
22902
ac833b4bb7ee moved some Drule operations to Thm (see more_thm.ML);
wenzelm
parents: 22709
diff changeset
   121
          map (HOLogic.dest_Trueprop o term_of o snd o Thm.dest_equals o cprop_of)
14166
1ed8f955727d Cleaned up the code.
skalberg
parents: 14164
diff changeset
   122
17336
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   123
        fun proc_single prop =
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   124
            let
29265
5b4247055bd7 moved old add_term_vars, add_term_frees etc. to structure OldTerm;
wenzelm
parents: 29006
diff changeset
   125
                val frees = OldTerm.term_frees prop
21858
05f57309170c avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
wenzelm
parents: 21434
diff changeset
   126
                val _ = forall (fn v => Sign.of_sort thy (type_of v,HOLogic.typeS)) frees
05f57309170c avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
wenzelm
parents: 21434
diff changeset
   127
                  orelse error "Specificaton: Only free variables of sort 'type' allowed"
35807
e4d1b5cbd429 added support for "specification" and "ax_specification" constructs to Nitpick
blanchet
parents: 35625
diff changeset
   128
                val prop_closed = close_form prop
17336
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   129
            in
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   130
                (prop_closed,frees)
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   131
            end
14166
1ed8f955727d Cleaned up the code.
skalberg
parents: 14164
diff changeset
   132
17336
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   133
        val props'' = map proc_single props'
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   134
        val frees = map snd props''
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   135
        val prop  = myfoldr HOLogic.mk_conj (map fst props'')
17895
6274b426594b moved helper lemma to HilbertChoice.thy;
wenzelm
parents: 17377
diff changeset
   136
        val cprop = cterm_of thy (HOLogic.mk_Trueprop prop)
14164
8c3fab596219 Allowed for splitting the specification over several lemmas.
skalberg
parents: 14122
diff changeset
   137
35845
e5980f0ad025 renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
wenzelm
parents: 35807
diff changeset
   138
        val (vmap, prop_thawed) = Type.varify_global [] prop
17336
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   139
        val thawed_prop_consts = collect_consts (prop_thawed,[])
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   140
        val (altcos,overloaded) = Library.split_list cos
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   141
        val (names,sconsts) = Library.split_list altcos
24707
dfeb98f84e93 Syntax.parse/check/read;
wenzelm
parents: 23591
diff changeset
   142
        val consts = map (Syntax.read_term_global thy) sconsts
21858
05f57309170c avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
wenzelm
parents: 21434
diff changeset
   143
        val _ = not (Library.exists (not o Term.is_Const) consts)
05f57309170c avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
wenzelm
parents: 21434
diff changeset
   144
          orelse error "Specification: Non-constant found as parameter"
14166
1ed8f955727d Cleaned up the code.
skalberg
parents: 14164
diff changeset
   145
17336
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   146
        fun proc_const c =
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   147
            let
35845
e5980f0ad025 renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
wenzelm
parents: 35807
diff changeset
   148
                val (_, c') = Type.varify_global [] c
17336
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   149
                val (cname,ctyp) = dest_Const c'
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   150
            in
33317
b4534348b8fd standardized filter/filter_out;
wenzelm
parents: 33245
diff changeset
   151
                case filter (fn t => let val (name,typ) = dest_Const t
19414
a21431e996bf Sign.typ_equiv;
wenzelm
parents: 18921
diff changeset
   152
                                     in name = cname andalso Sign.typ_equiv thy (typ, ctyp)
17336
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   153
                                     end) thawed_prop_consts of
26939
1035c89b4c02 moved global pretty/string_of functions from Sign to Syntax;
wenzelm
parents: 26928
diff changeset
   154
                    [] => error ("Specification: No suitable instances of constant \"" ^ Syntax.string_of_term_global thy c ^ "\" found")
35845
e5980f0ad025 renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
wenzelm
parents: 35807
diff changeset
   155
                  | [cf] => unvarify_global cf vmap
26939
1035c89b4c02 moved global pretty/string_of functions from Sign to Syntax;
wenzelm
parents: 26928
diff changeset
   156
                  | _ => error ("Specification: Several variations of \"" ^ Syntax.string_of_term_global thy c ^ "\" found (try applying explicit type constraints)")
17336
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   157
            end
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   158
        val proc_consts = map proc_const consts
33339
d41f77196338 eliminated some old folds;
wenzelm
parents: 33317
diff changeset
   159
        fun mk_exist c prop =
17336
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   160
            let
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   161
                val T = type_of c
30364
577edc39b501 moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents: 30280
diff changeset
   162
                val cname = Long_Name.base_name (fst (dest_Const c))
17336
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   163
                val vname = if Syntax.is_identifier cname
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   164
                            then cname
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   165
                            else "x"
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   166
            in
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   167
                HOLogic.exists_const T $ Abs(vname,T,Term.abstract_over (c,prop))
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   168
            end
33339
d41f77196338 eliminated some old folds;
wenzelm
parents: 33317
diff changeset
   169
        val ex_prop = fold_rev mk_exist proc_consts prop
17336
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   170
        val cnames = map (fst o dest_Const) proc_consts
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   171
        fun post_process (arg as (thy,thm)) =
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   172
            let
33245
65232054ffd0 eliminated some old folds;
wenzelm
parents: 32957
diff changeset
   173
                fun inst_all thy v thm =
17336
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   174
                    let
17895
6274b426594b moved helper lemma to HilbertChoice.thy;
wenzelm
parents: 17377
diff changeset
   175
                        val cv = cterm_of thy v
17336
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   176
                        val cT = ctyp_of_term cv
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   177
                        val spec' = instantiate' [SOME cT] [NONE,SOME cv] spec
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   178
                    in
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   179
                        thm RS spec'
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   180
                    end
22578
b0eb5652f210 removed obsolete sign_of/sign_of_thm;
wenzelm
parents: 22101
diff changeset
   181
                fun remove_alls frees thm =
33245
65232054ffd0 eliminated some old folds;
wenzelm
parents: 32957
diff changeset
   182
                    fold (inst_all (Thm.theory_of_thm thm)) frees thm
17336
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   183
                fun process_single ((name,atts),rew_imp,frees) args =
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   184
                    let
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   185
                        fun undo_imps thm =
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36618
diff changeset
   186
                            Thm.equal_elim (Thm.symmetric rew_imp) thm
14164
8c3fab596219 Allowed for splitting the specification over several lemmas.
skalberg
parents: 14122
diff changeset
   187
18358
0a733e11021a re-oriented some result tuples in PureThy
haftmann
parents: 17895
diff changeset
   188
                        fun add_final (arg as (thy, thm)) =
17336
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   189
                            if name = ""
18358
0a733e11021a re-oriented some result tuples in PureThy
haftmann
parents: 17895
diff changeset
   190
                            then arg |> Library.swap
32091
30e2ffbba718 proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
wenzelm
parents: 31723
diff changeset
   191
                            else (writeln ("  " ^ name ^ ": " ^ Display.string_of_thm_global thy thm);
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 38756
diff changeset
   192
                                  Global_Theory.store_thm (Binding.name name, thm) thy)
17336
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   193
                    in
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   194
                        args |> apsnd (remove_alls frees)
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   195
                             |> apsnd undo_imps
35021
c839a4c670c6 renamed old-style Drule.standard to Drule.export_without_context, to emphasize that this is in no way a standard operation;
wenzelm
parents: 33339
diff changeset
   196
                             |> apsnd Drule.export_without_context
35807
e4d1b5cbd429 added support for "specification" and "ax_specification" constructs to Nitpick
blanchet
parents: 35625
diff changeset
   197
                             |> Thm.theory_attributes (map (Attrib.attribute thy)
e4d1b5cbd429 added support for "specification" and "ax_specification" constructs to Nitpick
blanchet
parents: 35625
diff changeset
   198
                                                           (Attrib.internal (K Nitpick_Choice_Specs.add) :: atts))
17336
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   199
                             |> add_final
18358
0a733e11021a re-oriented some result tuples in PureThy
haftmann
parents: 17895
diff changeset
   200
                             |> Library.swap
17336
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   201
                    end
14164
8c3fab596219 Allowed for splitting the specification over several lemmas.
skalberg
parents: 14122
diff changeset
   202
17336
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   203
                fun process_all [proc_arg] args =
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   204
                    process_single proc_arg args
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   205
                  | process_all (proc_arg::rest) (thy,thm) =
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   206
                    let
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   207
                        val single_th = thm RS conjunct1
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   208
                        val rest_th   = thm RS conjunct2
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   209
                        val (thy',_)  = process_single proc_arg (thy,single_th)
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   210
                    in
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   211
                        process_all rest (thy',rest_th)
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   212
                    end
31723
f5cafe803b55 discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents: 30364
diff changeset
   213
                  | process_all [] _ = error "Choice_Specification.process_spec internal error"
17336
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   214
                val alt_names = map fst alt_props
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   215
                val _ = if exists (fn(name,_) => not (name = "")) alt_names
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   216
                        then writeln "specification"
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   217
                        else ()
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   218
            in
36618
7a0990473e03 specification goal: eliminated redundant Thm.legacy_freezeT -- the goal is properly declared and should always produce fixed types in the result;
wenzelm
parents: 36615
diff changeset
   219
                arg |> apsnd Thm.unvarify_global
17336
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   220
                    |> process_all (zip3 alt_names rew_imps frees)
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   221
            end
20903
905effde63d9 attribute: Context.mapping;
wenzelm
parents: 20548
diff changeset
   222
38756
d07959fabde6 renamed ProofContext.theory(_result) to ProofContext.background_theory(_result) to emphasize that this belongs to the infrastructure and is rarely appropriate in user-space tools;
wenzelm
parents: 38558
diff changeset
   223
      fun after_qed [[thm]] = ProofContext.background_theory (fn thy =>
d07959fabde6 renamed ProofContext.theory(_result) to ProofContext.background_theory(_result) to emphasize that this belongs to the infrastructure and is rarely appropriate in user-space tools;
wenzelm
parents: 38558
diff changeset
   224
        #1 (post_process (add_specification axiomatic (zip3 names cnames overloaded) (thy, thm))));
14115
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
   225
    in
20903
905effde63d9 attribute: Context.mapping;
wenzelm
parents: 20548
diff changeset
   226
      thy
36610
bafd82950e24 renamed ProofContext.init to ProofContext.init_global to emphasize that this is not the real thing;
wenzelm
parents: 36323
diff changeset
   227
      |> ProofContext.init_global
36618
7a0990473e03 specification goal: eliminated redundant Thm.legacy_freezeT -- the goal is properly declared and should always produce fixed types in the result;
wenzelm
parents: 36615
diff changeset
   228
      |> Variable.declare_term ex_prop
36323
655e2d74de3a modernized naming conventions of main Isar proof elements;
wenzelm
parents: 35897
diff changeset
   229
      |> Proof.theorem NONE after_qed [[(HOLogic.mk_Trueprop ex_prop, [])]]
20903
905effde63d9 attribute: Context.mapping;
wenzelm
parents: 20548
diff changeset
   230
    end;
14115
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
   231
17336
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   232
14115
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
   233
(* outer syntax *)
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
   234
36960
01594f816e3a prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents: 36954
diff changeset
   235
val opt_name = Scan.optional (Parse.name --| Parse.$$$ ":") ""
01594f816e3a prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents: 36954
diff changeset
   236
val opt_overloaded = Parse.opt_keyword "overloaded"
14116
63337d8e6e0f Added optional theorem names for the constant definitions added during
skalberg
parents: 14115
diff changeset
   237
14115
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
   238
val specification_decl =
36960
01594f816e3a prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents: 36954
diff changeset
   239
  Parse.$$$ "(" |-- Scan.repeat1 (opt_name -- Parse.term -- opt_overloaded) --| Parse.$$$ ")" --
01594f816e3a prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents: 36954
diff changeset
   240
          Scan.repeat1 ((Parse_Spec.opt_thm_name ":" >> apfst Binding.name_of) -- Parse.prop)
14115
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
   241
24867
e5b55d7be9bb simplified interfaces for outer syntax;
wenzelm
parents: 24707
diff changeset
   242
val _ =
36960
01594f816e3a prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents: 36954
diff changeset
   243
  Outer_Syntax.command "specification" "define constants by specification" Keyword.thy_goal
14164
8c3fab596219 Allowed for splitting the specification over several lemmas.
skalberg
parents: 14122
diff changeset
   244
    (specification_decl >> (fn (cos,alt_props) =>
17336
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   245
                               Toplevel.print o (Toplevel.theory_to_proof
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   246
                                                     (process_spec NONE cos alt_props))))
14222
1e61f23fd359 Added axiomatic specifications (ax_specification).
skalberg
parents: 14167
diff changeset
   247
1e61f23fd359 Added axiomatic specifications (ax_specification).
skalberg
parents: 14167
diff changeset
   248
val ax_specification_decl =
36960
01594f816e3a prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents: 36954
diff changeset
   249
    Parse.name --
01594f816e3a prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents: 36954
diff changeset
   250
    (Parse.$$$ "(" |-- Scan.repeat1 (opt_name -- Parse.term -- opt_overloaded) --| Parse.$$$ ")" --
01594f816e3a prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents: 36954
diff changeset
   251
           Scan.repeat1 ((Parse_Spec.opt_thm_name ":" >> apfst Binding.name_of) -- Parse.prop))
14115
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
   252
24867
e5b55d7be9bb simplified interfaces for outer syntax;
wenzelm
parents: 24707
diff changeset
   253
val _ =
36960
01594f816e3a prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents: 36954
diff changeset
   254
  Outer_Syntax.command "ax_specification" "define constants by specification" Keyword.thy_goal
14222
1e61f23fd359 Added axiomatic specifications (ax_specification).
skalberg
parents: 14167
diff changeset
   255
    (ax_specification_decl >> (fn (axname,(cos,alt_props)) =>
17336
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   256
                               Toplevel.print o (Toplevel.theory_to_proof
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   257
                                                     (process_spec (SOME axname) cos alt_props))))
14222
1e61f23fd359 Added axiomatic specifications (ax_specification).
skalberg
parents: 14167
diff changeset
   258
14115
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
   259
end