src/HOL/Tools/choice_specification.ML
author Simon Wimmer <wimmers@in.tum.de>
Thu, 18 Apr 2024 17:53:14 +0200
changeset 80137 0c51e0a6bc37
parent 79412 1c758cd8d5b2
permissions -rw-r--r--
sketch & explore: recover from duplicate fixed variables in Isar proofs
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
56270
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
    10
  val add_specification: (bstring * xstring * bool) list -> theory * thm -> theory * thm
14115
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
    11
end
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
    12
31723
f5cafe803b55 discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents: 30364
diff changeset
    13
structure Choice_Specification: CHOICE_SPECIFICATION =
14115
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
    14
struct
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
    15
56270
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
    16
local
46974
7ca3608146d8 eliminated odd 'finalconsts' / Theory.add_finals;
wenzelm
parents: 46961
diff changeset
    17
56270
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
    18
fun mk_definitional [] arg = arg
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
    19
  | mk_definitional ((thname, cname, covld) :: cos) (thy, thm) =
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 56895
diff changeset
    20
      (case HOLogic.dest_Trueprop (Thm.concl_of thm) of
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 61268
diff changeset
    21
        Const (\<^const_name>\<open>Ex\<close>, _) $ P =>
56270
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
    22
          let
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
    23
            val ctype = domain_type (type_of P)
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
    24
            val cname_full = Sign.intern_const thy cname
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
    25
            val cdefname =
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
    26
              if thname = ""
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
    27
              then Thm.def_name (Long_Name.base_name cname)
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
    28
              else thname
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
    29
            val def_eq =
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
    30
              Logic.mk_equals (Const (cname_full, ctype), HOLogic.choice_const ctype $ P)
79336
032a31db4c6f clarified signature: downgrade old-style Global_Theory.add_defs to Global_Theory.add_def without attributes;
wenzelm
parents: 79134
diff changeset
    31
            val (def_thm, thy') =
032a31db4c6f clarified signature: downgrade old-style Global_Theory.add_defs to Global_Theory.add_def without attributes;
wenzelm
parents: 79134
diff changeset
    32
              (if covld then Global_Theory.add_def_overloaded else  Global_Theory.add_def)
032a31db4c6f clarified signature: downgrade old-style Global_Theory.add_defs to Global_Theory.add_def without attributes;
wenzelm
parents: 79134
diff changeset
    33
                (Binding.name cdefname, def_eq) thy
032a31db4c6f clarified signature: downgrade old-style Global_Theory.add_defs to Global_Theory.add_def without attributes;
wenzelm
parents: 79134
diff changeset
    34
            val thm' = [thm, def_thm] MRS @{thm exE_some}
56270
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
    35
          in
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
    36
            mk_definitional cos (thy',thm')
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
    37
          end
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
    38
      | _ => raise THM ("Bad specification theorem", 0, [thm]))
14115
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
    39
14222
1e61f23fd359 Added axiomatic specifications (ax_specification).
skalberg
parents: 14167
diff changeset
    40
in
56270
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
    41
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
    42
fun add_specification cos =
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
    43
  mk_definitional cos #> apsnd Drule.export_without_context
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
    44
14222
1e61f23fd359 Added axiomatic specifications (ax_specification).
skalberg
parents: 14167
diff changeset
    45
end
1e61f23fd359 Added axiomatic specifications (ax_specification).
skalberg
parents: 14167
diff changeset
    46
18729
216e31270509 simplified type attribute;
wenzelm
parents: 18418
diff changeset
    47
14115
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
    48
(* Collect all intances of constants in term *)
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
    49
56270
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
    50
fun collect_consts (t $ u, tms) = collect_consts (u, collect_consts (t, tms))
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
    51
  | collect_consts (Abs (_, _, t), tms) = collect_consts (t, tms)
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
    52
  | collect_consts (tm as Const _, tms) = insert (op aconv) tm tms
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
    53
  | collect_consts (_, tms) = tms
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
    54
14115
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
    55
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
    56
(* Complementing Type.varify_global... *)
14115
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
    57
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
    58
fun unvarify_global t fmap =
56270
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
    59
  let
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
    60
    val fmap' = map Library.swap fmap
79134
5f0bbed1c606 misc tuning and clarification;
wenzelm
parents: 74266
diff changeset
    61
    fun unthaw v =
5f0bbed1c606 misc tuning and clarification;
wenzelm
parents: 74266
diff changeset
    62
      (case AList.lookup (op =) fmap' v of
79412
1c758cd8d5b2 minor performance tuning: proper Same.operation;
wenzelm
parents: 79336
diff changeset
    63
        NONE => raise Same.SAME
79134
5f0bbed1c606 misc tuning and clarification;
wenzelm
parents: 74266
diff changeset
    64
      | SOME w => TFree w)
56270
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
    65
  in map_types (map_type_tvar unthaw) t end
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
    66
14115
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
    67
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
    68
(* The syntactic meddling needed to setup add_specification for work *)
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
    69
56270
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
    70
fun close_form t =
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
    71
  fold_rev (fn (s, T) => fn t => HOLogic.mk_all (s, T, t))
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
    72
    (map dest_Free (Misc_Legacy.term_frees t)) t
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 50239
diff changeset
    73
56270
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
    74
fun zip3 [] [] [] = []
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
    75
  | zip3 (x::xs) (y::ys) (z::zs) = (x,y,z)::zip3 xs ys zs
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
    76
  | zip3 _ _ _ = raise Fail "Choice_Specification.process_spec internal error"
14167
3c29bba24aa4 Improved the error messages (slightly).
skalberg
parents: 14166
diff changeset
    77
56270
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
    78
fun myfoldr f [x] = x
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
    79
  | myfoldr f (x::xs) = f (x,myfoldr f xs)
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
    80
  | myfoldr f [] = raise Fail "Choice_Specification.process_spec internal error"
14164
8c3fab596219 Allowed for splitting the specification over several lemmas.
skalberg
parents: 14122
diff changeset
    81
56270
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
    82
fun process_spec cos alt_props thy =
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
    83
  let
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
    84
    val ctxt = Proof_Context.init_global thy
14166
1ed8f955727d Cleaned up the code.
skalberg
parents: 14164
diff changeset
    85
56270
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
    86
    val rew_imps = alt_props |>
59642
929984c529d3 clarified context;
wenzelm
parents: 59621
diff changeset
    87
      map (Object_Logic.atomize ctxt o Thm.cterm_of ctxt o Syntax.read_prop_global thy o snd)
56270
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
    88
    val props' = rew_imps |>
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 56895
diff changeset
    89
      map (HOLogic.dest_Trueprop o Thm.term_of o snd o Thm.dest_equals o Thm.cprop_of)
14166
1ed8f955727d Cleaned up the code.
skalberg
parents: 14164
diff changeset
    90
56270
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
    91
    fun proc_single prop =
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
    92
      let
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
    93
        val frees = Misc_Legacy.term_frees prop
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 61268
diff changeset
    94
        val _ = forall (fn v => Sign.of_sort thy (type_of v, \<^sort>\<open>type\<close>)) frees
56270
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
    95
          orelse error "Specificaton: Only free variables of sort 'type' allowed"
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
    96
        val prop_closed = close_form prop
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
    97
      in (prop_closed, frees) end
14164
8c3fab596219 Allowed for splitting the specification over several lemmas.
skalberg
parents: 14122
diff changeset
    98
56270
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
    99
    val props'' = map proc_single props'
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
   100
    val frees = map snd props''
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
   101
    val prop = myfoldr HOLogic.mk_conj (map fst props'')
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
   102
74266
612b7e0d6721 clarified signature;
wenzelm
parents: 74233
diff changeset
   103
    val (vmap, prop_thawed) = Type.varify_global TFrees.empty prop
56270
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
   104
    val thawed_prop_consts = collect_consts (prop_thawed, [])
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
   105
    val (altcos, overloaded) = split_list cos
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
   106
    val (names, sconsts) = split_list altcos
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
   107
    val consts = map (Syntax.read_term_global thy) sconsts
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
   108
    val _ = not (Library.exists (not o Term.is_Const) consts)
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
   109
      orelse error "Specification: Non-constant found as parameter"
14166
1ed8f955727d Cleaned up the code.
skalberg
parents: 14164
diff changeset
   110
56270
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
   111
    fun proc_const c =
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
   112
      let
74266
612b7e0d6721 clarified signature;
wenzelm
parents: 74233
diff changeset
   113
        val (_, c') = Type.varify_global TFrees.empty c
56270
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
   114
        val (cname,ctyp) = dest_Const c'
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
   115
      in
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
   116
        (case filter (fn t =>
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
   117
            let val (name, typ) = dest_Const t
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
   118
            in name = cname andalso Sign.typ_equiv thy (typ, ctyp)
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
   119
            end) thawed_prop_consts of
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
   120
          [] =>
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
   121
            error ("Specification: No suitable instances of constant \"" ^
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
   122
              Syntax.string_of_term_global thy c ^ "\" found")
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
   123
        | [cf] => unvarify_global cf vmap
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
   124
        | _ => error ("Specification: Several variations of \"" ^
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
   125
            Syntax.string_of_term_global thy c ^ "\" found (try applying explicit type constraints)"))
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
   126
      end
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
   127
    val proc_consts = map proc_const consts
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
   128
    fun mk_exist c prop =
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
   129
      let
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
   130
        val T = type_of c
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
   131
        val cname = Long_Name.base_name (fst (dest_Const c))
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
   132
        val vname = if Symbol_Pos.is_identifier cname then cname else "x"
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
   133
      in HOLogic.exists_const T $ Abs(vname,T,Term.abstract_over (c, prop)) end
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
   134
    val ex_prop = fold_rev mk_exist proc_consts prop
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
   135
    val cnames = map (fst o dest_Const) proc_consts
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
   136
    fun post_process (arg as (thy,thm)) =
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
   137
      let
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
   138
        fun inst_all thy v thm =
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
   139
          let
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59586
diff changeset
   140
            val cv = Thm.global_cterm_of thy v
59586
ddf6deaadfe8 clarified signature;
wenzelm
parents: 59582
diff changeset
   141
            val cT = Thm.ctyp_of_cterm cv
60801
7664e0916eec tuned signature;
wenzelm
parents: 60362
diff changeset
   142
            val spec' = Thm.instantiate' [SOME cT] [NONE, SOME cv] spec
56270
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
   143
          in thm RS spec' end
60362
befdc10ebb42 clarified context;
wenzelm
parents: 59936
diff changeset
   144
        fun remove_alls frees (thm, thy) = (fold (inst_all thy) frees thm, thy)
56270
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
   145
        fun process_single ((name, atts), rew_imp, frees) args =
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
   146
          let
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
   147
            fun undo_imps thm = Thm.equal_elim (Thm.symmetric rew_imp) thm
14164
8c3fab596219 Allowed for splitting the specification over several lemmas.
skalberg
parents: 14122
diff changeset
   148
56270
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
   149
            fun add_final (thm, thy) =
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
   150
              if name = ""
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
   151
              then (thm, thy)
61268
abe08fb15a12 moved remaining display.ML to more_thm.ML;
wenzelm
parents: 60825
diff changeset
   152
              else (writeln ("  " ^ name ^ ": " ^ Thm.string_of_thm_global thy thm);
56270
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
   153
                Global_Theory.store_thm (Binding.name name, thm) thy)
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
   154
          in
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
   155
            swap args
60362
befdc10ebb42 clarified context;
wenzelm
parents: 59936
diff changeset
   156
            |> remove_alls frees
56270
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
   157
            |> apfst undo_imps
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
   158
            |> apfst Drule.export_without_context
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
   159
            |-> Thm.theory_attributes
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
   160
              (map (Attrib.attribute_cmd_global thy)
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
   161
                (@{attributes [nitpick_choice_spec]} @ atts))
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
   162
            |> add_final
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
   163
            |> swap
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
   164
          end
14164
8c3fab596219 Allowed for splitting the specification over several lemmas.
skalberg
parents: 14122
diff changeset
   165
56270
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
   166
        fun process_all [proc_arg] args =
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
   167
            process_single proc_arg args
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
   168
          | process_all (proc_arg::rest) (thy,thm) =
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
   169
              let
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
   170
                val single_th = thm RS conjunct1
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
   171
                val rest_th = thm RS conjunct2
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
   172
                val (thy', _) = process_single proc_arg (thy, single_th)
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
   173
              in process_all rest (thy', rest_th) end
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
   174
          | process_all [] _ = raise Fail "Choice_Specification.process_spec internal error"
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
   175
        val alt_names = map fst alt_props
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
   176
        val _ =
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
   177
          if exists (fn (name, _) => name <> "") alt_names
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
   178
          then writeln "specification"
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
   179
          else ()
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
   180
      in
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
   181
        arg
60825
bacfb7c45d81 more explicit context;
wenzelm
parents: 60801
diff changeset
   182
        |> apsnd (Thm.unvarify_global thy)
56270
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
   183
        |> process_all (zip3 alt_names rew_imps frees)
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
   184
      end
20903
905effde63d9 attribute: Context.mapping;
wenzelm
parents: 20548
diff changeset
   185
56270
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
   186
    fun after_qed [[thm]] = Proof_Context.background_theory (fn thy =>
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
   187
      #1 (post_process (add_specification (zip3 names cnames overloaded) (thy, thm))));
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
   188
  in
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
   189
    thy
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
   190
    |> Proof_Context.init_global
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
   191
    |> Variable.declare_term ex_prop
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
   192
    |> Proof.theorem NONE after_qed [[(HOLogic.mk_Trueprop ex_prop, [])]]
ce9c7a527c4b removed unused 'ax_specification', to give 'specification' a chance to become localized;
wenzelm
parents: 56254
diff changeset
   193
  end
14115
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
   194
17336
c05f72cff368 tuned IsarThy.theorem_i;
wenzelm
parents: 17057
diff changeset
   195
14115
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
   196
(* outer syntax *)
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
   197
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 61268
diff changeset
   198
val opt_name = Scan.optional (Parse.name --| \<^keyword>\<open>:\<close>) ""
36960
01594f816e3a prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents: 36954
diff changeset
   199
val opt_overloaded = Parse.opt_keyword "overloaded"
14116
63337d8e6e0f Added optional theorem names for the constant definitions added during
skalberg
parents: 14115
diff changeset
   200
46961
5c6955f487e5 outer syntax command definitions based on formal command_spec derived from theory header declarations;
wenzelm
parents: 46949
diff changeset
   201
val _ =
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 61268
diff changeset
   202
  Outer_Syntax.command \<^command_keyword>\<open>specification\<close> "define constants by specification"
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 61268
diff changeset
   203
    (\<^keyword>\<open>(\<close> |-- Scan.repeat1 (opt_name -- Parse.term -- opt_overloaded) --| \<^keyword>\<open>)\<close> --
46961
5c6955f487e5 outer syntax command definitions based on formal command_spec derived from theory header declarations;
wenzelm
parents: 46949
diff changeset
   204
      Scan.repeat1 ((Parse_Spec.opt_thm_name ":" >> apfst Binding.name_of) -- Parse.prop)
56895
f058120aaad4 discontinued Toplevel.print flag -- print uniformly according to Keyword.is_printed;
wenzelm
parents: 56270
diff changeset
   205
      >> (fn (cos, alt_props) => Toplevel.theory_to_proof (process_spec cos alt_props)))
14222
1e61f23fd359 Added axiomatic specifications (ax_specification).
skalberg
parents: 14167
diff changeset
   206
14115
65ec3f73d00b Added package for definition by specification.
skalberg
parents:
diff changeset
   207
end