src/Pure/Isar/isar_thy.ML
author wenzelm
Thu, 19 Nov 1998 11:49:09 +0100
changeset 5938 fe7640933a47
parent 5925 669d0bc621e1
child 5959 7af99477751b
permissions -rw-r--r--
match_bind: 'as' patterns; statements: 'is' patterns;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
5830
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
     1
(*  Title:      Pure/Isar/isar_thy.ML
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
     2
    ID:         $Id$
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
     3
    Author:     Markus Wenzel, TU Muenchen
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
     4
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
     5
Derived theory operations.
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
     6
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
     7
TODO:
5925
669d0bc621e1 tuned comments;
wenzelm
parents: 5915
diff changeset
     8
  - pure_thy.ML: add_constdefs (atomic!);
5830
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
     9
  - 'methods' section (proof macros, ML method defs) (!?);
5882
c8096461f5c8 renamed tac / etac to refine / then_refine;
wenzelm
parents: 5830
diff changeset
    10
  - next_block: ProofHistory open / close (!?);
5830
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
    11
*)
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
    12
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
    13
signature ISAR_THY =
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
    14
sig
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
    15
  val add_axioms: ((bstring * string) * Args.src list) list -> theory -> theory
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
    16
  val add_defs: ((bstring * string) * Args.src list) list -> theory -> theory
5915
66f4bde4c6e7 added have_theorems, have_lemmas, have_facts;
wenzelm
parents: 5882
diff changeset
    17
  val have_theorems: (bstring * Args.src list) * (xstring * Args.src list) list
66f4bde4c6e7 added have_theorems, have_lemmas, have_facts;
wenzelm
parents: 5882
diff changeset
    18
    -> theory -> theory
66f4bde4c6e7 added have_theorems, have_lemmas, have_facts;
wenzelm
parents: 5882
diff changeset
    19
  val have_lemmas: (bstring * Args.src list) * (xstring * Args.src list) list
66f4bde4c6e7 added have_theorems, have_lemmas, have_facts;
wenzelm
parents: 5882
diff changeset
    20
    -> theory -> theory
66f4bde4c6e7 added have_theorems, have_lemmas, have_facts;
wenzelm
parents: 5882
diff changeset
    21
  val have_facts: (string * Args.src list) * (string * Args.src list) list
66f4bde4c6e7 added have_theorems, have_lemmas, have_facts;
wenzelm
parents: 5882
diff changeset
    22
    -> ProofHistory.T -> ProofHistory.T
5830
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
    23
  val fix: (string * string option) list -> ProofHistory.T -> ProofHistory.T
5938
fe7640933a47 match_bind: 'as' patterns;
wenzelm
parents: 5925
diff changeset
    24
  val match_bind: (string list * string) list -> ProofHistory.T -> ProofHistory.T
fe7640933a47 match_bind: 'as' patterns;
wenzelm
parents: 5925
diff changeset
    25
  val theorem: string -> Args.src list -> string * string list -> theory -> ProofHistory.T
fe7640933a47 match_bind: 'as' patterns;
wenzelm
parents: 5925
diff changeset
    26
  val lemma: string -> Args.src list -> string * string list -> theory -> ProofHistory.T
fe7640933a47 match_bind: 'as' patterns;
wenzelm
parents: 5925
diff changeset
    27
  val assume: string -> Args.src list -> (string * string list) list
fe7640933a47 match_bind: 'as' patterns;
wenzelm
parents: 5925
diff changeset
    28
    -> ProofHistory.T -> ProofHistory.T
fe7640933a47 match_bind: 'as' patterns;
wenzelm
parents: 5925
diff changeset
    29
  val show: string -> Args.src list -> string * string list -> ProofHistory.T -> ProofHistory.T
fe7640933a47 match_bind: 'as' patterns;
wenzelm
parents: 5925
diff changeset
    30
  val have: string -> Args.src list -> string * string list -> ProofHistory.T -> ProofHistory.T
5830
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
    31
  val chain: ProofHistory.T -> ProofHistory.T
5915
66f4bde4c6e7 added have_theorems, have_lemmas, have_facts;
wenzelm
parents: 5882
diff changeset
    32
  val from_facts: (string * Args.src list) list -> ProofHistory.T -> ProofHistory.T
5830
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
    33
  val begin_block: ProofHistory.T -> ProofHistory.T
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
    34
  val next_block: ProofHistory.T -> ProofHistory.T
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
    35
  val end_block: ProofHistory.T -> ProofHistory.T
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
    36
  val qed_with: bstring option * Args.src list option -> ProofHistory.T -> theory
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
    37
  val qed: ProofHistory.T -> theory
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
    38
  val kill_proof: ProofHistory.T -> theory
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
    39
  val tac: Method.text -> ProofHistory.T -> ProofHistory.T
5882
c8096461f5c8 renamed tac / etac to refine / then_refine;
wenzelm
parents: 5830
diff changeset
    40
  val then_tac: Method.text -> ProofHistory.T -> ProofHistory.T
5830
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
    41
  val proof: Method.text option -> ProofHistory.T -> ProofHistory.T
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
    42
  val terminal_proof: Method.text -> ProofHistory.T -> ProofHistory.T
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
    43
  val trivial_proof: ProofHistory.T -> ProofHistory.T
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
    44
  val default_proof: ProofHistory.T -> ProofHistory.T
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
    45
  val use_mltext: string -> theory option -> theory option
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
    46
  val use_mltext_theory: string -> theory -> theory
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
    47
  val use_setup: string -> theory -> theory
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
    48
  val parse_ast_translation: string -> theory -> theory
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
    49
  val parse_translation: string -> theory -> theory
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
    50
  val print_translation: string -> theory -> theory
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
    51
  val typed_print_translation: string -> theory -> theory
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
    52
  val print_ast_translation: string -> theory -> theory
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
    53
  val token_translation: string -> theory -> theory
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
    54
  val add_oracle: bstring * string -> theory -> theory
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
    55
  val the_theory: string -> unit -> theory
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
    56
  val begin_theory: string * string list -> unit -> theory
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
    57
  val end_theory: theory -> unit
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
    58
end;
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
    59
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
    60
structure IsarThy: ISAR_THY =
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
    61
struct
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
    62
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
    63
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
    64
(** derived theory and proof operations **)
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
    65
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
    66
(* axioms and defs *)
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
    67
5915
66f4bde4c6e7 added have_theorems, have_lemmas, have_facts;
wenzelm
parents: 5882
diff changeset
    68
fun add_axms f args thy =
66f4bde4c6e7 added have_theorems, have_lemmas, have_facts;
wenzelm
parents: 5882
diff changeset
    69
  f (map (fn (x, srcs) => (x, map (Attrib.global_attribute thy) srcs)) args) thy;
66f4bde4c6e7 added have_theorems, have_lemmas, have_facts;
wenzelm
parents: 5882
diff changeset
    70
66f4bde4c6e7 added have_theorems, have_lemmas, have_facts;
wenzelm
parents: 5882
diff changeset
    71
val add_axioms = add_axms PureThy.add_axioms;
66f4bde4c6e7 added have_theorems, have_lemmas, have_facts;
wenzelm
parents: 5882
diff changeset
    72
val add_defs = add_axms PureThy.add_defs;
66f4bde4c6e7 added have_theorems, have_lemmas, have_facts;
wenzelm
parents: 5882
diff changeset
    73
66f4bde4c6e7 added have_theorems, have_lemmas, have_facts;
wenzelm
parents: 5882
diff changeset
    74
66f4bde4c6e7 added have_theorems, have_lemmas, have_facts;
wenzelm
parents: 5882
diff changeset
    75
(* theorems *)
66f4bde4c6e7 added have_theorems, have_lemmas, have_facts;
wenzelm
parents: 5882
diff changeset
    76
66f4bde4c6e7 added have_theorems, have_lemmas, have_facts;
wenzelm
parents: 5882
diff changeset
    77
fun gen_have_thmss get attrib f ((name, more_srcs), th_srcs) x =
66f4bde4c6e7 added have_theorems, have_lemmas, have_facts;
wenzelm
parents: 5882
diff changeset
    78
  f name (map (attrib x) more_srcs)
66f4bde4c6e7 added have_theorems, have_lemmas, have_facts;
wenzelm
parents: 5882
diff changeset
    79
    (map (fn (s, srcs) => (get x s, map (attrib x) srcs)) th_srcs) x;
66f4bde4c6e7 added have_theorems, have_lemmas, have_facts;
wenzelm
parents: 5882
diff changeset
    80
5830
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
    81
5915
66f4bde4c6e7 added have_theorems, have_lemmas, have_facts;
wenzelm
parents: 5882
diff changeset
    82
val have_theorems =
66f4bde4c6e7 added have_theorems, have_lemmas, have_facts;
wenzelm
parents: 5882
diff changeset
    83
  #1 oo gen_have_thmss PureThy.get_tthms Attrib.global_attribute PureThy.have_tthmss;
66f4bde4c6e7 added have_theorems, have_lemmas, have_facts;
wenzelm
parents: 5882
diff changeset
    84
66f4bde4c6e7 added have_theorems, have_lemmas, have_facts;
wenzelm
parents: 5882
diff changeset
    85
val have_lemmas =
66f4bde4c6e7 added have_theorems, have_lemmas, have_facts;
wenzelm
parents: 5882
diff changeset
    86
  #1 oo gen_have_thmss PureThy.get_tthms Attrib.global_attribute
66f4bde4c6e7 added have_theorems, have_lemmas, have_facts;
wenzelm
parents: 5882
diff changeset
    87
    (fn name => fn more_atts => PureThy.have_tthmss name (more_atts @ [Attribute.tag_lemma]));
66f4bde4c6e7 added have_theorems, have_lemmas, have_facts;
wenzelm
parents: 5882
diff changeset
    88
66f4bde4c6e7 added have_theorems, have_lemmas, have_facts;
wenzelm
parents: 5882
diff changeset
    89
66f4bde4c6e7 added have_theorems, have_lemmas, have_facts;
wenzelm
parents: 5882
diff changeset
    90
val have_thmss =
66f4bde4c6e7 added have_theorems, have_lemmas, have_facts;
wenzelm
parents: 5882
diff changeset
    91
  gen_have_thmss (ProofContext.get_tthms o Proof.context_of)
66f4bde4c6e7 added have_theorems, have_lemmas, have_facts;
wenzelm
parents: 5882
diff changeset
    92
    (Attrib.local_attribute o Proof.theory_of) Proof.have_tthmss;
66f4bde4c6e7 added have_theorems, have_lemmas, have_facts;
wenzelm
parents: 5882
diff changeset
    93
66f4bde4c6e7 added have_theorems, have_lemmas, have_facts;
wenzelm
parents: 5882
diff changeset
    94
val have_facts = ProofHistory.apply o have_thmss;
66f4bde4c6e7 added have_theorems, have_lemmas, have_facts;
wenzelm
parents: 5882
diff changeset
    95
val from_facts = ProofHistory.apply o (Proof.chain oo curry have_thmss ("", []));
5830
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
    96
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
    97
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
    98
(* context *)
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
    99
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   100
val fix = ProofHistory.apply o Proof.fix;
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   101
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   102
val match_bind = ProofHistory.apply o Proof.match_bind;
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   103
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   104
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   105
(* statements *)
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   106
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   107
fun global_statement f name tags s thy =
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   108
  ProofHistory.init (f name (map (Attrib.global_attribute thy) tags) s thy);
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   109
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   110
fun local_statement f name tags s = ProofHistory.apply_open (fn state =>
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   111
  f name (map (Attrib.local_attribute (Proof.theory_of state)) tags) s state);
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   112
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   113
val theorem = global_statement Proof.theorem;
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   114
val lemma = global_statement Proof.lemma;
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   115
val assume = local_statement Proof.assume;
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   116
val show = local_statement Proof.show;
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   117
val have = local_statement Proof.have;
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   118
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   119
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   120
(* forward chaining *)
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   121
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   122
val chain = ProofHistory.apply Proof.chain;
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   123
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   124
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   125
(* end proof *)
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   126
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   127
fun qed_with (alt_name, alt_tags) prf =
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   128
  let
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   129
    val state = ProofHistory.current prf;
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   130
    val thy = Proof.theory_of state;
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   131
    val alt_atts = apsome (map (Attrib.global_attribute thy)) alt_tags;
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   132
    val (thy', (kind, name, tthm)) = Method.qed alt_name alt_atts state;
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   133
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   134
    val prt_result = Pretty.block
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   135
      [Pretty.str (kind ^ " " ^ name ^ ":"), Pretty.fbrk, Attribute.pretty_tthm tthm];
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   136
  in Pretty.writeln prt_result; thy end;
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   137
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   138
val qed = qed_with (None, None);
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   139
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   140
val kill_proof = Proof.theory_of o ProofHistory.current;
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   141
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   142
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   143
(* blocks *)
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   144
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   145
val begin_block = ProofHistory.apply_open Proof.begin_block;
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   146
val next_block = ProofHistory.apply Proof.next_block;
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   147
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   148
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   149
(* backward steps *)
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   150
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   151
val tac = ProofHistory.applys o Method.tac;
5882
c8096461f5c8 renamed tac / etac to refine / then_refine;
wenzelm
parents: 5830
diff changeset
   152
val then_tac = ProofHistory.applys o Method.then_tac;
5830
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   153
val proof = ProofHistory.applys o Method.proof;
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   154
val end_block = ProofHistory.applys_close Method.end_block;
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   155
val terminal_proof = ProofHistory.applys_close o Method.terminal_proof;
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   156
val trivial_proof = ProofHistory.applys_close Method.trivial_proof;
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   157
val default_proof = ProofHistory.applys_close Method.default_proof;
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   158
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   159
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   160
(* use ML text *)
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   161
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   162
fun use_mltext txt opt_thy =
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   163
  let
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   164
    val save_context = Context.get_context ();
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   165
    val _ = Context.set_context opt_thy;
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   166
    val opt_exn = (transform_error (use_text false) txt; None) handle exn => Some exn;
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   167
    val opt_thy' = Context.get_context ();
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   168
  in
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   169
    Context.set_context save_context;
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   170
    (case opt_exn of
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   171
      None => opt_thy'
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   172
    | Some exn => raise exn)
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   173
  end;
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   174
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   175
fun use_mltext_theory txt thy =
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   176
  (case use_mltext txt (Some thy) of
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   177
    Some thy' => thy'
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   178
  | None => error "Missing result ML theory context");
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   179
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   180
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   181
fun use_context txt = use_mltext_theory ("Context.>> (" ^ txt ^ ");");
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   182
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   183
fun use_let name body txt =
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   184
  use_context ("let val " ^ name ^ " = " ^ txt ^ " in\n" ^ body ^ " end");
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   185
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   186
val use_setup =
5915
66f4bde4c6e7 added have_theorems, have_lemmas, have_facts;
wenzelm
parents: 5882
diff changeset
   187
  use_let "setup: (theory -> theory) list" "Library.apply setup";
5830
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   188
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   189
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   190
(* translation functions *)
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   191
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   192
val parse_ast_translation =
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   193
  use_let "parse_ast_translation: (string * (Syntax.ast list -> Syntax.ast)) list"
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   194
    "Theory.add_trfuns (parse_ast_translation, [], [], [])";
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   195
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   196
val parse_translation =
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   197
  use_let "parse_translation: (string * (term list -> term)) list"
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   198
    "Theory.add_trfuns ([], parse_translation, [], [])";
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   199
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   200
val print_translation =
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   201
  use_let "print_translation: (string * (term list -> term)) list"
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   202
    "Theory.add_trfuns ([], [], print_translation, [])";
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   203
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   204
val print_ast_translation =
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   205
  use_let "print_ast_translation: (string * (Syntax.ast list -> Syntax.ast)) list"
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   206
    "Theory.add_trfuns ([], [], [], print_ast_translation)";
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   207
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   208
val typed_print_translation =
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   209
  use_let "typed_print_translation: (string * (bool -> typ -> term list -> term)) list"
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   210
    "Theory.add_trfunsT typed_print_translation";
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   211
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   212
val token_translation =
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   213
  use_let "token_translation: (string * string * (string -> string * int)) list"
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   214
    "Theory.add_tokentrfuns token_translation";
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   215
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   216
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   217
(* add_oracle *)
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   218
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   219
fun add_oracle (name, txt) =
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   220
  use_let
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   221
    "oracle: bstring * (Sign.sg * Object.T -> term)"
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   222
    "Theory.add_oracle oracle"
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   223
    ("(" ^ quote name ^ ", " ^ txt ^ ")");
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   224
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   225
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   226
(* theory init and exit *)
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   227
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   228
fun the_theory name () = ThyInfo.theory name;
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   229
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   230
fun begin_theory (name, names) () =
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   231
  PureThy.begin_theory name (map ThyInfo.theory names);
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   232
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   233
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   234
fun end_theory thy =
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   235
  let val thy' = PureThy.end_theory thy in
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   236
    transform_error ThyInfo.store_theory thy'
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   237
      handle exn => raise PureThy.ROLLBACK (thy', Some exn)     (* FIXME !!?? *)
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   238
  end;
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   239
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   240
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   241
end;