src/Pure/Isar/local_theory.ML
author wenzelm
Wed, 06 Sep 2006 22:48:36 +0200
changeset 20487 6ac7a4fc32a0
parent 20252 bad805d0456b
child 20778 f39c733f2a7e
permissions -rw-r--r--
read_instantiate: declare names of TVars as well (temporary workaround for no-freeze feature of type inference);
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
18951
4f5f6c632127 type local_theory = Proof.context;
wenzelm
parents: 18876
diff changeset
     8
type local_theory = Proof.context;
4f5f6c632127 type local_theory = Proof.context;
wenzelm
parents: 18876
diff changeset
     9
18744
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
    10
signature LOCAL_THEORY =
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
    11
sig
18951
4f5f6c632127 type local_theory = Proof.context;
wenzelm
parents: 18876
diff changeset
    12
  val params_of: local_theory -> (string * typ) list
4f5f6c632127 type local_theory = Proof.context;
wenzelm
parents: 18876
diff changeset
    13
  val hyps_of: local_theory -> term list
19913
3f844845ecb8 standard: simultaneous facts;
wenzelm
parents: 19680
diff changeset
    14
  val standard: local_theory -> thm list -> thm list
18951
4f5f6c632127 type local_theory = Proof.context;
wenzelm
parents: 18876
diff changeset
    15
  val pretty_consts: local_theory -> (string * typ -> bool) -> (string * typ) list -> Pretty.T
4f5f6c632127 type local_theory = Proof.context;
wenzelm
parents: 18876
diff changeset
    16
  val quiet_mode: bool ref
4f5f6c632127 type local_theory = Proof.context;
wenzelm
parents: 18876
diff changeset
    17
  val print_consts: local_theory -> (string * typ -> bool) -> (string * typ) list -> unit
4f5f6c632127 type local_theory = Proof.context;
wenzelm
parents: 18876
diff changeset
    18
  val theory: (theory -> theory) -> local_theory -> local_theory
4f5f6c632127 type local_theory = Proof.context;
wenzelm
parents: 18876
diff changeset
    19
  val theory_result: (theory -> 'a * theory) -> local_theory -> 'a * local_theory
4f5f6c632127 type local_theory = Proof.context;
wenzelm
parents: 18876
diff changeset
    20
  val init: xstring option -> theory -> local_theory
4f5f6c632127 type local_theory = Proof.context;
wenzelm
parents: 18876
diff changeset
    21
  val init_i: string option -> theory -> local_theory
19017
5f06c37043e4 added restore;
wenzelm
parents: 18951
diff changeset
    22
  val restore: local_theory -> local_theory
18951
4f5f6c632127 type local_theory = Proof.context;
wenzelm
parents: 18876
diff changeset
    23
  val exit: local_theory -> local_theory * theory
19059
b4ca3100e818 init/exit no longer change the theory (no naming);
wenzelm
parents: 19032
diff changeset
    24
  val restore_naming: local_theory -> local_theory
b4ca3100e818 init/exit no longer change the theory (no naming);
wenzelm
parents: 19032
diff changeset
    25
  val naming: local_theory -> local_theory
b4ca3100e818 init/exit no longer change the theory (no naming);
wenzelm
parents: 19032
diff changeset
    26
  val mapping: xstring option -> (local_theory -> local_theory) -> theory -> local_theory * theory
19680
6a34b1b1f540 export generic term_syntax;
wenzelm
parents: 19661
diff changeset
    27
  val term_syntax: (Context.generic -> Context.generic) -> local_theory -> local_theory
19661
baab85f25e0e added syntax interface;
wenzelm
parents: 19630
diff changeset
    28
  val syntax: string * bool -> (string * mixfix) list -> local_theory -> local_theory
baab85f25e0e added syntax interface;
wenzelm
parents: 19630
diff changeset
    29
  val abbrevs: string * bool -> ((bstring * mixfix) * term) list -> local_theory -> local_theory
18951
4f5f6c632127 type local_theory = Proof.context;
wenzelm
parents: 18876
diff changeset
    30
  val consts: ((bstring * typ) * mixfix) list -> local_theory -> term list * local_theory
18876
ddb6803da197 added consts_retricted;
wenzelm
parents: 18823
diff changeset
    31
  val consts_restricted: (string * typ -> bool) ->
18951
4f5f6c632127 type local_theory = Proof.context;
wenzelm
parents: 18876
diff changeset
    32
    ((bstring * typ) * mixfix) list -> local_theory -> term list * local_theory
4f5f6c632127 type local_theory = Proof.context;
wenzelm
parents: 18876
diff changeset
    33
  val axioms: ((bstring * Attrib.src list) * term list) list -> local_theory ->
4f5f6c632127 type local_theory = Proof.context;
wenzelm
parents: 18876
diff changeset
    34
    (bstring * thm list) list * local_theory
4f5f6c632127 type local_theory = Proof.context;
wenzelm
parents: 18876
diff changeset
    35
  val axioms_finish: (local_theory -> thm -> thm) ->
4f5f6c632127 type local_theory = Proof.context;
wenzelm
parents: 18876
diff changeset
    36
    ((bstring * Attrib.src list) * term list) list -> local_theory ->
4f5f6c632127 type local_theory = Proof.context;
wenzelm
parents: 18876
diff changeset
    37
    (bstring * thm list) list * local_theory
18767
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
    38
  val notes: string -> ((bstring * Attrib.src list) * (thm list * Attrib.src list) list) list ->
18951
4f5f6c632127 type local_theory = Proof.context;
wenzelm
parents: 18876
diff changeset
    39
    local_theory -> (bstring * thm list) list * local_theory
4f5f6c632127 type local_theory = Proof.context;
wenzelm
parents: 18876
diff changeset
    40
  val note: (bstring * Attrib.src list) * thm list -> local_theory ->
4f5f6c632127 type local_theory = Proof.context;
wenzelm
parents: 18876
diff changeset
    41
    (bstring * thm list) * local_theory
18823
916c493b7f0c added print_consts;
wenzelm
parents: 18805
diff changeset
    42
  val def: (bstring * mixfix) * ((bstring * Attrib.src list) * term) ->
18951
4f5f6c632127 type local_theory = Proof.context;
wenzelm
parents: 18876
diff changeset
    43
    local_theory -> (term * (bstring * thm)) * local_theory
4f5f6c632127 type local_theory = Proof.context;
wenzelm
parents: 18876
diff changeset
    44
  val def_finish: (local_theory -> term -> thm -> thm) ->
18823
916c493b7f0c added print_consts;
wenzelm
parents: 18805
diff changeset
    45
    (bstring * mixfix) * ((bstring * Attrib.src list) * term) ->
18951
4f5f6c632127 type local_theory = Proof.context;
wenzelm
parents: 18876
diff changeset
    46
    local_theory -> (term * (bstring * thm)) * local_theory
18744
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
    47
end;
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
    48
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
    49
structure LocalTheory: LOCAL_THEORY =
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
    50
struct
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
    51
19059
b4ca3100e818 init/exit no longer change the theory (no naming);
wenzelm
parents: 19032
diff changeset
    52
18767
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
    53
(** local context data **)
18744
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
    54
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
    55
structure Data = ProofDataFun
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
    56
(
18876
ddb6803da197 added consts_retricted;
wenzelm
parents: 18823
diff changeset
    57
  val name = "Pure/local_theory";
18744
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
    58
  type T =
20032
2087e5634598 removed obsolete locale view;
wenzelm
parents: 19991
diff changeset
    59
   {locale: (string * Proof.context) option,
18744
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
    60
    params: (string * typ) list,
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
    61
    hyps: term list,
19059
b4ca3100e818 init/exit no longer change the theory (no naming);
wenzelm
parents: 19032
diff changeset
    62
    restore_naming: theory -> theory};
b4ca3100e818 init/exit no longer change the theory (no naming);
wenzelm
parents: 19032
diff changeset
    63
  fun init thy = {locale = NONE, params = [], hyps = [], restore_naming = I};
18744
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
    64
  fun print _ _ = ();
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
    65
);
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
    66
val _ = Context.add_setup Data.init;
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
    67
19059
b4ca3100e818 init/exit no longer change the theory (no naming);
wenzelm
parents: 19032
diff changeset
    68
fun map_locale f = Data.map (fn {locale, params, hyps, restore_naming} =>
20032
2087e5634598 removed obsolete locale view;
wenzelm
parents: 19991
diff changeset
    69
  {locale = Option.map (fn (loc, ctxt) => (loc, f ctxt)) locale,
19059
b4ca3100e818 init/exit no longer change the theory (no naming);
wenzelm
parents: 19032
diff changeset
    70
    params = params, hyps = hyps, restore_naming = restore_naming});
18805
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
    71
18767
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
    72
val locale_of = #locale o Data.get;
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
    73
val params_of = #params o Data.get;
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
    74
val hyps_of = #hyps o Data.get;
18805
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
    75
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
    76
fun standard ctxt =
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
    77
  (case #locale (Data.get ctxt) of
20252
bad805d0456b tuned comment;
wenzelm
parents: 20243
diff changeset
    78
    NONE => map Drule.standard   (* FIXME !? *)
20032
2087e5634598 removed obsolete locale view;
wenzelm
parents: 19991
diff changeset
    79
  | SOME (_, loc_ctxt) => ProofContext.export_standard ctxt loc_ctxt);
18744
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
    80
18767
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
    81
18823
916c493b7f0c added print_consts;
wenzelm
parents: 18805
diff changeset
    82
(* print consts *)
18767
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
    83
18951
4f5f6c632127 type local_theory = Proof.context;
wenzelm
parents: 18876
diff changeset
    84
val quiet_mode = ref false;
4f5f6c632127 type local_theory = Proof.context;
wenzelm
parents: 18876
diff changeset
    85
18767
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
    86
local
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
    87
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
    88
fun pretty_var ctxt (x, T) =
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
    89
  Pretty.block [Pretty.str x, Pretty.str " ::", Pretty.brk 1,
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
    90
    Pretty.quote (ProofContext.pretty_typ ctxt T)];
18744
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
    91
18767
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
    92
fun pretty_vars ctxt kind vs = Pretty.big_list kind (map (pretty_var ctxt) vs);
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
    93
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
    94
in
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
    95
18876
ddb6803da197 added consts_retricted;
wenzelm
parents: 18823
diff changeset
    96
fun pretty_consts ctxt pred cs =
ddb6803da197 added consts_retricted;
wenzelm
parents: 18823
diff changeset
    97
  (case filter pred (params_of ctxt) of
18767
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
    98
    [] => pretty_vars ctxt "constants" cs
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
    99
  | ps => Pretty.chunks [pretty_vars ctxt "parameters" ps, pretty_vars ctxt "constants" cs]);
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
   100
18876
ddb6803da197 added consts_retricted;
wenzelm
parents: 18823
diff changeset
   101
fun print_consts _ _ [] = ()
18951
4f5f6c632127 type local_theory = Proof.context;
wenzelm
parents: 18876
diff changeset
   102
  | print_consts ctxt pred cs =
4f5f6c632127 type local_theory = Proof.context;
wenzelm
parents: 18876
diff changeset
   103
      if ! quiet_mode then () else Pretty.writeln (pretty_consts ctxt pred cs);
18823
916c493b7f0c added print_consts;
wenzelm
parents: 18805
diff changeset
   104
18767
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
   105
end;
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
   106
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
   107
18805
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
   108
(* theory/locale substructures *)
18781
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   109
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   110
fun theory_result f ctxt =
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   111
  let val (res, thy') = f (ProofContext.theory_of ctxt)
19032
wenzelm
parents: 19017
diff changeset
   112
  in (res, ctxt |> map_locale (ProofContext.transfer thy') |> ProofContext.transfer thy') end;
wenzelm
parents: 19017
diff changeset
   113
wenzelm
parents: 19017
diff changeset
   114
fun theory f ctxt = #2 (theory_result (f #> pair ()) ctxt);
18805
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
   115
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
   116
fun locale_result f ctxt =
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
   117
  (case locale_of ctxt of
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
   118
    NONE => error "Local theory: no locale context"
20032
2087e5634598 removed obsolete locale view;
wenzelm
parents: 19991
diff changeset
   119
  | SOME (_, loc_ctxt) =>
18805
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
   120
      let
20032
2087e5634598 removed obsolete locale view;
wenzelm
parents: 19991
diff changeset
   121
        val (res, loc_ctxt') = f loc_ctxt;
18805
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
   122
        val thy' = ProofContext.theory_of loc_ctxt';
19032
wenzelm
parents: 19017
diff changeset
   123
      in (res, ctxt |> map_locale (K loc_ctxt') |> ProofContext.transfer thy') end);
18805
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
   124
19017
5f06c37043e4 added restore;
wenzelm
parents: 18951
diff changeset
   125
fun locale f ctxt =
5f06c37043e4 added restore;
wenzelm
parents: 18951
diff changeset
   126
  if is_none (locale_of ctxt) then ctxt
5f06c37043e4 added restore;
wenzelm
parents: 18951
diff changeset
   127
  else #2 (locale_result (f #> pair ()) ctxt);
18805
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
   128
19017
5f06c37043e4 added restore;
wenzelm
parents: 18951
diff changeset
   129
5f06c37043e4 added restore;
wenzelm
parents: 18951
diff changeset
   130
(* init -- restore -- exit *)
18781
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   131
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   132
fun init_i NONE thy = ProofContext.init thy
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   133
  | init_i (SOME loc) thy =
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   134
      thy
20032
2087e5634598 removed obsolete locale view;
wenzelm
parents: 19991
diff changeset
   135
      |> Locale.init loc
2087e5634598 removed obsolete locale view;
wenzelm
parents: 19991
diff changeset
   136
      |> ProofContext.inferred_fixes
2087e5634598 removed obsolete locale view;
wenzelm
parents: 19991
diff changeset
   137
      |> (fn (params, ctxt) => Data.put
2087e5634598 removed obsolete locale view;
wenzelm
parents: 19991
diff changeset
   138
          {locale = SOME (loc, ctxt),
18781
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   139
           params = params,
20236
54e15870444b "moved basic assumption operations from structure ProofContext to Assumption;"
wenzelm
parents: 20032
diff changeset
   140
           hyps = Assumption.assms_of ctxt,
19059
b4ca3100e818 init/exit no longer change the theory (no naming);
wenzelm
parents: 19032
diff changeset
   141
           restore_naming = Sign.restore_naming thy} ctxt);
18781
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   142
18805
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
   143
fun init loc thy = init_i (Option.map (Locale.intern thy) loc) thy;
18781
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   144
19017
5f06c37043e4 added restore;
wenzelm
parents: 18951
diff changeset
   145
fun restore ctxt =
5f06c37043e4 added restore;
wenzelm
parents: 18951
diff changeset
   146
  (case locale_of ctxt of
5f06c37043e4 added restore;
wenzelm
parents: 18951
diff changeset
   147
    NONE => ctxt
20032
2087e5634598 removed obsolete locale view;
wenzelm
parents: 19991
diff changeset
   148
  | SOME (_, loc_ctxt) => loc_ctxt |> Data.put (Data.get ctxt));
19017
5f06c37043e4 added restore;
wenzelm
parents: 18951
diff changeset
   149
19059
b4ca3100e818 init/exit no longer change the theory (no naming);
wenzelm
parents: 19032
diff changeset
   150
fun exit ctxt = (ctxt, ProofContext.theory_of ctxt);
b4ca3100e818 init/exit no longer change the theory (no naming);
wenzelm
parents: 19032
diff changeset
   151
b4ca3100e818 init/exit no longer change the theory (no naming);
wenzelm
parents: 19032
diff changeset
   152
b4ca3100e818 init/exit no longer change the theory (no naming);
wenzelm
parents: 19032
diff changeset
   153
(* local naming *)
b4ca3100e818 init/exit no longer change the theory (no naming);
wenzelm
parents: 19032
diff changeset
   154
b4ca3100e818 init/exit no longer change the theory (no naming);
wenzelm
parents: 19032
diff changeset
   155
fun naming ctxt =
b4ca3100e818 init/exit no longer change the theory (no naming);
wenzelm
parents: 19032
diff changeset
   156
  (case locale_of ctxt of
b4ca3100e818 init/exit no longer change the theory (no naming);
wenzelm
parents: 19032
diff changeset
   157
    NONE => ctxt
b4ca3100e818 init/exit no longer change the theory (no naming);
wenzelm
parents: 19032
diff changeset
   158
  | SOME (loc, _) => ctxt |> theory (Sign.sticky_prefix (Sign.base_name loc)));
b4ca3100e818 init/exit no longer change the theory (no naming);
wenzelm
parents: 19032
diff changeset
   159
b4ca3100e818 init/exit no longer change the theory (no naming);
wenzelm
parents: 19032
diff changeset
   160
fun restore_naming ctxt = theory (#restore_naming (Data.get ctxt)) ctxt;
b4ca3100e818 init/exit no longer change the theory (no naming);
wenzelm
parents: 19032
diff changeset
   161
b4ca3100e818 init/exit no longer change the theory (no naming);
wenzelm
parents: 19032
diff changeset
   162
fun mapping loc f =
b4ca3100e818 init/exit no longer change the theory (no naming);
wenzelm
parents: 19032
diff changeset
   163
  init loc #> naming #> f #> restore_naming #> exit;
18781
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   164
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   165
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   166
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   167
(** local theory **)
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   168
19661
baab85f25e0e added syntax interface;
wenzelm
parents: 19630
diff changeset
   169
(* term syntax *)
baab85f25e0e added syntax interface;
wenzelm
parents: 19630
diff changeset
   170
19680
6a34b1b1f540 export generic term_syntax;
wenzelm
parents: 19661
diff changeset
   171
fun term_syntax f ctxt = ctxt |>
19661
baab85f25e0e added syntax interface;
wenzelm
parents: 19630
diff changeset
   172
  (case locale_of ctxt of
19680
6a34b1b1f540 export generic term_syntax;
wenzelm
parents: 19661
diff changeset
   173
    NONE => theory (Context.theory_map f)
19661
baab85f25e0e added syntax interface;
wenzelm
parents: 19630
diff changeset
   174
  | SOME (loc, _) =>
19680
6a34b1b1f540 export generic term_syntax;
wenzelm
parents: 19661
diff changeset
   175
      locale (Locale.add_term_syntax loc (Context.proof_map f)) #>
6a34b1b1f540 export generic term_syntax;
wenzelm
parents: 19661
diff changeset
   176
      Context.proof_map f);
19661
baab85f25e0e added syntax interface;
wenzelm
parents: 19630
diff changeset
   177
19680
6a34b1b1f540 export generic term_syntax;
wenzelm
parents: 19661
diff changeset
   178
fun syntax x y =
6a34b1b1f540 export generic term_syntax;
wenzelm
parents: 19661
diff changeset
   179
  term_syntax (Context.mapping (Sign.add_const_syntax x y) (ProofContext.add_const_syntax x y));
19661
baab85f25e0e added syntax interface;
wenzelm
parents: 19630
diff changeset
   180
19680
6a34b1b1f540 export generic term_syntax;
wenzelm
parents: 19661
diff changeset
   181
fun abbrevs x y =
6a34b1b1f540 export generic term_syntax;
wenzelm
parents: 19661
diff changeset
   182
  term_syntax (Context.mapping (Sign.add_abbrevs x y) (ProofContext.add_abbrevs x y));
19661
baab85f25e0e added syntax interface;
wenzelm
parents: 19630
diff changeset
   183
baab85f25e0e added syntax interface;
wenzelm
parents: 19630
diff changeset
   184
18767
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
   185
(* consts *)
18744
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
   186
18876
ddb6803da197 added consts_retricted;
wenzelm
parents: 18823
diff changeset
   187
fun consts_restricted pred decls ctxt =
18744
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
   188
  let
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
   189
    val thy = ProofContext.theory_of ctxt;
18876
ddb6803da197 added consts_retricted;
wenzelm
parents: 18823
diff changeset
   190
    val params = filter pred (params_of ctxt);
18744
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
   191
    val ps = map Free params;
18805
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
   192
    val Ps = map #2 params;
19661
baab85f25e0e added syntax interface;
wenzelm
parents: 19630
diff changeset
   193
    val abbrs = decls |> map (fn ((c, T), mx) =>
baab85f25e0e added syntax interface;
wenzelm
parents: 19630
diff changeset
   194
      ((c, mx), Term.list_comb (Const (Sign.full_name thy c, Ps ---> T), ps)));
18744
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
   195
  in
19017
5f06c37043e4 added restore;
wenzelm
parents: 18951
diff changeset
   196
    ctxt |>
5f06c37043e4 added restore;
wenzelm
parents: 18951
diff changeset
   197
    (case locale_of ctxt of
19661
baab85f25e0e added syntax interface;
wenzelm
parents: 19630
diff changeset
   198
      NONE => theory (Sign.add_consts_authentic (map (fn ((c, T), mx) => (c, T, mx)) decls))
19017
5f06c37043e4 added restore;
wenzelm
parents: 18951
diff changeset
   199
    | SOME (loc, _) =>
5f06c37043e4 added restore;
wenzelm
parents: 18951
diff changeset
   200
        theory (Sign.add_consts_i (map (fn ((c, T), _) => (c, Ps ---> T, NoSyn)) decls)) #>
19661
baab85f25e0e added syntax interface;
wenzelm
parents: 19630
diff changeset
   201
        abbrevs Syntax.default_mode abbrs)
baab85f25e0e added syntax interface;
wenzelm
parents: 19630
diff changeset
   202
    |> pair (map #2 abbrs)
18767
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
   203
  end;
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
   204
18876
ddb6803da197 added consts_retricted;
wenzelm
parents: 18823
diff changeset
   205
val consts = consts_restricted (K true);
ddb6803da197 added consts_retricted;
wenzelm
parents: 18823
diff changeset
   206
18767
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
   207
18805
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
   208
(* fact definitions *)
18767
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
   209
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
   210
fun notes kind facts ctxt =
18805
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
   211
  let
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
   212
    val attrib = Attrib.attribute_i (ProofContext.theory_of ctxt);
19913
3f844845ecb8 standard: simultaneous facts;
wenzelm
parents: 19680
diff changeset
   213
    val facts' = map (apsnd (map (apfst (standard ctxt)))) facts;  (* FIXME burrow standard *)
18805
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
   214
  in
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
   215
    ctxt |>
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
   216
    (case locale_of ctxt of
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
   217
      NONE => theory_result (PureThy.note_thmss_i kind (Attrib.map_facts attrib facts'))
19017
5f06c37043e4 added restore;
wenzelm
parents: 18951
diff changeset
   218
    | SOME (loc, _) =>
20032
2087e5634598 removed obsolete locale view;
wenzelm
parents: 19991
diff changeset
   219
        locale_result (apfst #1 o Locale.add_thmss kind loc facts'))
18805
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
   220
    ||> (#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
   221
  end;
18767
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
   222
18781
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   223
fun note (a, ths) ctxt =
18805
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
   224
  ctxt |> notes PureThy.theoremK [(a, [(ths, [])])] |>> hd;
18767
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
   225
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
   226
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
   227
(* axioms *)
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
   228
18781
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   229
local
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   230
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   231
fun add_axiom hyps (name, prop) thy =
18767
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
   232
  let
19517
wenzelm
parents: 19370
diff changeset
   233
    val name' = if name = "" then "axiom_" ^ serial_string () else name;
18767
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
   234
    val frees = rev ([] |> Term.add_frees prop |> fold Term.add_frees hyps);
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
   235
    val prop' = Term.list_all_free (frees, Logic.list_implies (hyps, prop));
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
   236
    val thy' = thy |> Theory.add_axioms_i [(name', prop')];
18781
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   237
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   238
    val cert = Thm.cterm_of thy';
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   239
    fun all_polymorphic (x, T) th =
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   240
      let val var = cert (Var ((x, 0), #1 (Logic.dest_all (Thm.prop_of th))))
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   241
      in ((var, cert (Free (x, T))), Thm.forall_elim var th) end;
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   242
    fun implies_polymorphic hyp th = Thm.assume (cert hyp) COMP th;
18767
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
   243
    val th =
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
   244
      Thm.get_axiom_i thy' (Sign.full_name thy' name')
18781
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   245
      |> fold_map all_polymorphic frees |-> Drule.cterm_instantiate
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   246
      |> fold implies_polymorphic hyps
18767
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
   247
  in (th, thy') end;
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
   248
18781
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   249
in
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   250
18805
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
   251
fun axioms_finish finish = fold_map (fn ((name, atts), props) =>
20243
8b64a1ea9b1b renamed ProofContext.fix_frees to Variable.fix_frees;
wenzelm
parents: 20236
diff changeset
   252
  fold Variable.fix_frees props
18805
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
   253
  #> (fn ctxt => ctxt
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
   254
    |> 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
   255
    |-> (fn ths => note ((name, atts), map (finish ctxt) ths))));
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
   256
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
   257
val axioms = axioms_finish (K I);
18744
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
   258
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
   259
end;
18781
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   260
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   261
18805
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
   262
(* constant definitions *)
18781
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   263
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   264
local
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   265
18805
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
   266
fun add_def (name, prop) =
19630
d370c3f5d3b2 Theory.add_defs(_i): added unchecked flag;
wenzelm
parents: 19517
diff changeset
   267
  Theory.add_defs_i false false [(name, prop)] #> (fn thy =>
18805
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
   268
    let
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
   269
      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
   270
      val cert = Thm.cterm_of thy;
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
   271
      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
   272
        (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
   273
    in (Drule.cterm_instantiate subst th, thy) end);
18781
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   274
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   275
in
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   276
18805
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
   277
fun def_finish finish (var, spec) ctxt =
18781
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   278
  let
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   279
    val (x, mx) = var;
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   280
    val ((name, atts), rhs) = spec;
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   281
    val name' = if name = "" then Thm.def_name x else name;
18876
ddb6803da197 added consts_retricted;
wenzelm
parents: 18823
diff changeset
   282
    val rhs_frees = Term.add_frees rhs [];
18781
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   283
  in
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   284
    ctxt
18876
ddb6803da197 added consts_retricted;
wenzelm
parents: 18823
diff changeset
   285
    |> consts_restricted (member (op =) rhs_frees) [((x, Term.fastype_of rhs), mx)]
18781
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   286
    |-> (fn [lhs] =>
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   287
      theory_result (add_def (name', Logic.mk_equals (lhs, rhs)))
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   288
      #-> (fn th => fn ctxt' => note ((name', atts), [finish ctxt' lhs th]) ctxt')
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   289
      #> apfst (fn (b, [th]) => (lhs, (b, th))))
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   290
  end;
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   291
18805
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
   292
val def = def_finish (K (K I));
18863b33c831 improved 'notes', including proper treatment of locale results;
wenzelm
parents: 18781
diff changeset
   293
18781
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   294
end;
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   295
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   296
end;