src/Pure/Isar/local_theory.ML
author wenzelm
Fri, 27 Jan 2006 19:03:11 +0100
changeset 18805 18863b33c831
parent 18781 ea3b5e22eab5
child 18823 916c493b7f0c
permissions -rw-r--r--
improved 'notes', including proper treatment of locale results; tuned;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
18744
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
     1
(*  Title:      Pure/Isar/local_theory.ML
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
     2
    ID:         $Id$
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
     3
    Author:     Makarius
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
     4
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
     5
Local theory operations, with optional target locale.
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
     6
*)
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
     7
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
     8
signature LOCAL_THEORY =
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
     9
sig
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
    10
  val params_of: Proof.context -> (string * typ) list
18767
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
    11
  val hyps_of: Proof.context -> term list
18781
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
    12
  val standard: Proof.context -> thm -> thm
18767
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
    13
  val pretty_consts: Proof.context -> (string * typ) list -> Pretty.T
18781
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
    14
  val theory: (theory -> theory) -> Proof.context -> Proof.context
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
    15
  val theory_result: (theory -> 'a * theory) -> Proof.context -> 'a * Proof.context
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
    16
  val init: xstring option -> theory -> Proof.context
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
    17
  val init_i: string option -> theory -> Proof.context
18805
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
    18
  val exit: Proof.context -> Proof.context * theory
18767
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
    19
  val consts: ((string * typ) * mixfix) list -> Proof.context -> term list * Proof.context
18781
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
    20
  val axioms: ((string * Attrib.src list) * term list) list -> Proof.context ->
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
    21
    (bstring * thm list) list * Proof.context
18805
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
    22
  val axioms_finish: (Proof.context -> thm -> thm) ->
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
    23
    ((string * Attrib.src list) * term list) list -> Proof.context ->
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
    24
    (bstring * thm list) list * Proof.context
18767
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
    25
  val notes: string -> ((bstring * Attrib.src list) * (thm list * Attrib.src list) list) list ->
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
    26
    Proof.context -> (bstring * thm list) list * Proof.context
18781
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
    27
  val note: (bstring * Attrib.src list) * thm list -> Proof.context ->
18767
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
    28
    (bstring * thm list) * Proof.context
18781
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
    29
  val def: (string * mixfix) * ((string * Attrib.src list) * term) ->
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
    30
    Proof.context -> (term * (bstring * thm)) * Proof.context
18805
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
    31
  val def_finish: (Proof.context -> term -> thm -> thm) ->
18781
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
    32
    (string * mixfix) * ((string * Attrib.src list) * term) ->
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
    33
    Proof.context -> (term * (bstring * thm)) * Proof.context
18744
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
    34
end;
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
    35
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
    36
structure LocalTheory: LOCAL_THEORY =
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
    37
struct
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
    38
18767
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
    39
(** local context data **)
18744
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
    40
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
    41
structure Data = ProofDataFun
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
    42
(
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
    43
  val name = "Isar/local_theory";
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
    44
  type T =
18805
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
    45
   {locale: (string * (cterm list * Proof.context)) option,
18744
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
    46
    params: (string * typ) list,
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
    47
    hyps: term list,
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
    48
    exit: theory -> theory};
18805
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
    49
  fun init thy = {locale = NONE, params = [], hyps = [], exit = I};
18744
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
    50
  fun print _ _ = ();
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
    51
);
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
    52
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
    53
val _ = Context.add_setup Data.init;
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
    54
18805
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
    55
fun map_locale f = Data.map (fn {locale, params, hyps, exit} =>
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
    56
  {locale = Option.map (fn (loc, (view, ctxt)) => (loc, (view, f ctxt))) locale,
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
    57
    params = params, hyps = hyps, exit = exit});
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
    58
18767
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
    59
val locale_of = #locale o Data.get;
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
    60
val params_of = #params o Data.get;
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
    61
val hyps_of = #hyps o Data.get;
18805
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
    62
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
    63
fun standard ctxt =
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
    64
  (case #locale (Data.get ctxt) of
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
    65
    NONE => Drule.standard
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
    66
  | SOME (_, (_, loc_ctxt)) => ProofContext.export_standard ctxt loc_ctxt);
18744
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
    67
18767
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
    68
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
    69
(* pretty_consts *)
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
    70
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
    71
local
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
    72
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
    73
fun pretty_var ctxt (x, T) =
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
    74
  Pretty.block [Pretty.str x, Pretty.str " ::", Pretty.brk 1,
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
    75
    Pretty.quote (ProofContext.pretty_typ ctxt T)];
18744
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
    76
18767
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
    77
fun pretty_vars ctxt kind vs = Pretty.big_list kind (map (pretty_var ctxt) vs);
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
    78
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
    79
in
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
    80
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
    81
fun pretty_consts ctxt cs =
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
    82
  (case params_of ctxt of
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
    83
    [] => pretty_vars ctxt "constants" cs
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
    84
  | ps => Pretty.chunks [pretty_vars ctxt "parameters" ps, pretty_vars ctxt "constants" cs]);
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
    85
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
    86
end;
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
    87
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
    88
18805
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
    89
(* theory/locale substructures *)
18781
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
    90
18805
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
    91
fun transfer thy =
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
    92
  ProofContext.transfer thy #> map_locale (ProofContext.transfer thy);
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
    93
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
    94
fun theory f ctxt = transfer (f (ProofContext.theory_of ctxt)) ctxt;
18781
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
    95
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
    96
fun theory_result f ctxt =
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
    97
  let val (res, thy') = f (ProofContext.theory_of ctxt)
18805
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
    98
  in (res, transfer thy' ctxt) end;
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
    99
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
   100
fun locale_result f ctxt =
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
   101
  (case locale_of ctxt of
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
   102
    NONE => error "Local theory: no locale context"
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
   103
  | SOME (loc, (view, loc_ctxt)) =>
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
   104
      let
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
   105
        val (res, loc_ctxt') = f loc_ctxt;
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
   106
        val thy' = ProofContext.theory_of loc_ctxt';
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
   107
      in (res, ctxt |> map_locale (K loc_ctxt') |> transfer thy') end);
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
   108
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
   109
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
   110
(* init -- exit *)
18781
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   111
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   112
fun init_i NONE thy = ProofContext.init thy
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   113
  | init_i (SOME loc) thy =
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   114
      thy
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   115
      |> Locale.init loc
18805
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
   116
      ||>> ProofContext.inferred_fixes
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
   117
      |> (fn ((view, params), ctxt) => Data.put
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
   118
          {locale = SOME (loc, (view, ctxt)),
18781
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   119
           params = params,
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   120
           hyps = ProofContext.assms_of ctxt,
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   121
           exit = Sign.restore_naming thy} ctxt)
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   122
      |> theory (Sign.add_path (Sign.base_name loc) #> Sign.no_base_names);
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   123
18805
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
   124
fun init loc thy = init_i (Option.map (Locale.intern thy) loc) thy;
18781
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   125
18805
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
   126
fun exit ctxt = (ctxt, #exit (Data.get ctxt) (ProofContext.theory_of ctxt));
18781
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   127
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   128
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   129
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   130
(** local theory **)
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   131
18767
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
   132
(* consts *)
18744
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
   133
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
   134
fun consts decls ctxt =
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
   135
  let
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
   136
    val thy = ProofContext.theory_of ctxt;
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
   137
    val params = params_of ctxt;
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
   138
    val ps = map Free params;
18805
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
   139
    val Ps = map #2 params;
18744
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
   140
  in
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
   141
    ctxt
18781
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   142
    |> theory (Sign.add_consts_i (decls |> map (fn ((c, T), mx) => (c, Ps ---> T, mx))))
18767
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
   143
    |> pair (decls |> map (fn ((c, T), _) =>
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
   144
      Term.list_comb (Const (Sign.full_name thy c, Ps ---> T), ps)))
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
   145
  end;
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
   146
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
   147
18805
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
   148
(* fact definitions *)
18767
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
   149
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
   150
fun notes kind facts ctxt =
18805
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
   151
  let
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
   152
    val attrib = Attrib.attribute_i (ProofContext.theory_of ctxt);
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
   153
    val facts' = map (apsnd (map (apfst (map (standard ctxt))))) facts;
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
   154
  in
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
   155
    ctxt |>
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
   156
    (case locale_of ctxt of
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
   157
      NONE => theory_result (PureThy.note_thmss_i kind (Attrib.map_facts attrib facts'))
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
   158
    | SOME (loc, view_ctxt) =>
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
   159
        locale_result (K (apfst #1 (Locale.add_thmss kind loc facts' view_ctxt))))
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
   160
    ||> (#2 o ProofContext.note_thmss_i (Attrib.map_facts attrib facts))
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
   161
  end;
18767
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
   162
18781
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   163
fun note (a, ths) ctxt =
18805
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
   164
  ctxt |> notes PureThy.theoremK [(a, [(ths, [])])] |>> hd;
18767
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
   165
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
   166
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
   167
(* axioms *)
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
   168
18781
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   169
local
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   170
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   171
fun add_axiom hyps (name, prop) thy =
18767
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
   172
  let
18781
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   173
    val name' = if name = "" then "axiom_" ^ string_of_int (serial ()) else name;
18767
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
   174
    val frees = rev ([] |> Term.add_frees prop |> fold Term.add_frees hyps);
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
   175
    val prop' = Term.list_all_free (frees, Logic.list_implies (hyps, prop));
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
   176
    val thy' = thy |> Theory.add_axioms_i [(name', prop')];
18781
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   177
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   178
    val cert = Thm.cterm_of thy';
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   179
    fun all_polymorphic (x, T) th =
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   180
      let val var = cert (Var ((x, 0), #1 (Logic.dest_all (Thm.prop_of th))))
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   181
      in ((var, cert (Free (x, T))), Thm.forall_elim var th) end;
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   182
    fun implies_polymorphic hyp th = Thm.assume (cert hyp) COMP th;
18767
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
   183
    val th =
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
   184
      Thm.get_axiom_i thy' (Sign.full_name thy' name')
18781
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   185
      |> fold_map all_polymorphic frees |-> Drule.cterm_instantiate
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   186
      |> fold implies_polymorphic hyps
18767
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
   187
  in (th, thy') end;
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
   188
18781
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   189
in
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   190
18805
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
   191
fun axioms_finish finish = fold_map (fn ((name, atts), props) =>
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
   192
  fold ProofContext.fix_frees props
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
   193
  #> (fn ctxt => ctxt
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
   194
    |> theory_result (fold_map (add_axiom (hyps_of ctxt)) (PureThy.name_multi name props))
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
   195
    |-> (fn ths => note ((name, atts), map (finish ctxt) ths))));
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
   196
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
   197
val axioms = axioms_finish (K I);
18744
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
   198
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
   199
end;
18781
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   200
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   201
18805
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
   202
(* constant definitions *)
18781
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   203
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   204
local
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   205
18805
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
   206
fun add_def (name, prop) =
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
   207
  Theory.add_defs_i false [(name, prop)] #> (fn thy =>
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
   208
    let
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
   209
      val th = Thm.get_axiom_i thy (Sign.full_name thy name);
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
   210
      val cert = Thm.cterm_of thy;
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
   211
      val subst = map2 (fn var => fn free => (cert (Var var), cert (Free free)))
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
   212
        (Term.add_vars (Thm.prop_of th) []) (Term.add_frees prop []);
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
   213
    in (Drule.cterm_instantiate subst th, thy) end);
18781
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   214
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   215
in
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   216
18805
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
   217
fun def_finish finish (var, spec) ctxt =
18781
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   218
  let
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   219
    val (x, mx) = var;
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   220
    val ((name, atts), rhs) = spec;
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   221
    val name' = if name = "" then Thm.def_name x else name;
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   222
  in
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   223
    ctxt
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   224
    |> consts [((x, Term.fastype_of rhs), mx)]
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   225
    |-> (fn [lhs] =>
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   226
      theory_result (add_def (name', Logic.mk_equals (lhs, rhs)))
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   227
      #-> (fn th => fn ctxt' => note ((name', atts), [finish ctxt' lhs th]) ctxt')
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   228
      #> apfst (fn (b, [th]) => (lhs, (b, th))))
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   229
  end;
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   230
18805
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
   231
val def = def_finish (K (K I));
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
   232
18781
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   233
end;
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   234
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   235
end;