src/Pure/Isar/isar_thy.ML
author wenzelm
Fri, 23 Apr 1999 16:33:23 +0200
changeset 6501 392333eb31cb
parent 6483 3e5d450c2b31
child 6531 8064ed198068
permissions -rw-r--r--
added thus, hence;
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
6371
8469852acbc0 added '_i' versions;
wenzelm
parents: 6354
diff changeset
     5
Pure/Isar derived theory operations.
5830
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
     6
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
     7
TODO:
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
     8
  - 'methods' section (proof macros, ML method defs) (!?);
5882
c8096461f5c8 renamed tac / etac to refine / then_refine;
wenzelm
parents: 5830
diff changeset
     9
  - next_block: ProofHistory open / close (!?);
5830
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
    10
*)
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
    11
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
    12
signature ISAR_THY =
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
    13
sig
5959
7af99477751b add_text, add_chapter etc.: dummy;
wenzelm
parents: 5938
diff changeset
    14
  val add_text: string -> theory -> theory
6354
a4c75cbd2fbf add_title;
wenzelm
parents: 6331
diff changeset
    15
  val add_title: string -> string -> string -> theory -> theory
5959
7af99477751b add_text, add_chapter etc.: dummy;
wenzelm
parents: 5938
diff changeset
    16
  val add_chapter: string -> theory -> theory
7af99477751b add_text, add_chapter etc.: dummy;
wenzelm
parents: 5938
diff changeset
    17
  val add_section: string -> theory -> theory
7af99477751b add_text, add_chapter etc.: dummy;
wenzelm
parents: 5938
diff changeset
    18
  val add_subsection: string -> theory -> theory
7af99477751b add_text, add_chapter etc.: dummy;
wenzelm
parents: 5938
diff changeset
    19
  val add_subsubsection: string -> theory -> theory
5830
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
    20
  val add_axioms: ((bstring * string) * Args.src list) list -> theory -> theory
6371
8469852acbc0 added '_i' versions;
wenzelm
parents: 6354
diff changeset
    21
  val add_axioms_i: ((bstring * term) * theory attribute list) list -> theory -> theory
5830
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
    22
  val add_defs: ((bstring * string) * Args.src list) list -> theory -> theory
6371
8469852acbc0 added '_i' versions;
wenzelm
parents: 6354
diff changeset
    23
  val add_defs_i: ((bstring * term) * theory attribute list) list -> theory -> theory
8469852acbc0 added '_i' versions;
wenzelm
parents: 6354
diff changeset
    24
  val add_constdefs: ((bstring * string * mixfix) * string) list -> theory -> theory
8469852acbc0 added '_i' versions;
wenzelm
parents: 6354
diff changeset
    25
  val add_constdefs_i: ((bstring * typ * mixfix) * term) list -> theory -> theory
8469852acbc0 added '_i' versions;
wenzelm
parents: 6354
diff changeset
    26
  val apply_theorems: (xstring * Args.src list) list -> theory -> theory * thm list
8469852acbc0 added '_i' versions;
wenzelm
parents: 6354
diff changeset
    27
  val apply_theorems_i: (thm * theory attribute list) list -> theory -> theory * thm list
5915
66f4bde4c6e7 added have_theorems, have_lemmas, have_facts;
wenzelm
parents: 5882
diff changeset
    28
  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
    29
    -> theory -> theory
6371
8469852acbc0 added '_i' versions;
wenzelm
parents: 6354
diff changeset
    30
  val have_theorems_i: (bstring * theory attribute list) * (thm * theory attribute list) list
8469852acbc0 added '_i' versions;
wenzelm
parents: 6354
diff changeset
    31
    -> theory -> theory
5915
66f4bde4c6e7 added have_theorems, have_lemmas, have_facts;
wenzelm
parents: 5882
diff changeset
    32
  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
    33
    -> theory -> theory
6371
8469852acbc0 added '_i' versions;
wenzelm
parents: 6354
diff changeset
    34
  val have_lemmas_i: (bstring * theory attribute list) * (thm * theory attribute list) list
8469852acbc0 added '_i' versions;
wenzelm
parents: 6354
diff changeset
    35
    -> theory -> theory
5915
66f4bde4c6e7 added have_theorems, have_lemmas, have_facts;
wenzelm
parents: 5882
diff changeset
    36
  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
    37
    -> ProofHistory.T -> ProofHistory.T
6371
8469852acbc0 added '_i' versions;
wenzelm
parents: 6354
diff changeset
    38
  val have_facts_i: (string * Proof.context attribute list) *
8469852acbc0 added '_i' versions;
wenzelm
parents: 6354
diff changeset
    39
    (thm * Proof.context attribute list) list -> ProofHistory.T -> ProofHistory.T
8469852acbc0 added '_i' versions;
wenzelm
parents: 6354
diff changeset
    40
  val from_facts: (string * Args.src list) list -> ProofHistory.T -> ProofHistory.T
8469852acbc0 added '_i' versions;
wenzelm
parents: 6354
diff changeset
    41
  val from_facts_i: (thm * Proof.context attribute list) list -> ProofHistory.T -> ProofHistory.T
8469852acbc0 added '_i' versions;
wenzelm
parents: 6354
diff changeset
    42
  val chain: ProofHistory.T -> ProofHistory.T
5830
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
    43
  val fix: (string * string option) list -> ProofHistory.T -> ProofHistory.T
6371
8469852acbc0 added '_i' versions;
wenzelm
parents: 6354
diff changeset
    44
  val fix_i: (string * typ) list -> ProofHistory.T -> ProofHistory.T
5938
fe7640933a47 match_bind: 'as' patterns;
wenzelm
parents: 5925
diff changeset
    45
  val match_bind: (string list * string) list -> ProofHistory.T -> ProofHistory.T
6371
8469852acbc0 added '_i' versions;
wenzelm
parents: 6354
diff changeset
    46
  val match_bind_i: (term list * term) list -> ProofHistory.T -> ProofHistory.T
5938
fe7640933a47 match_bind: 'as' patterns;
wenzelm
parents: 5925
diff changeset
    47
  val theorem: string -> Args.src list -> string * string list -> theory -> ProofHistory.T
6371
8469852acbc0 added '_i' versions;
wenzelm
parents: 6354
diff changeset
    48
  val theorem_i: bstring -> theory attribute list -> term * term list -> theory -> ProofHistory.T
5938
fe7640933a47 match_bind: 'as' patterns;
wenzelm
parents: 5925
diff changeset
    49
  val lemma: string -> Args.src list -> string * string list -> theory -> ProofHistory.T
6371
8469852acbc0 added '_i' versions;
wenzelm
parents: 6354
diff changeset
    50
  val lemma_i: bstring -> theory attribute list -> term * term list -> theory -> ProofHistory.T
5938
fe7640933a47 match_bind: 'as' patterns;
wenzelm
parents: 5925
diff changeset
    51
  val assume: string -> Args.src list -> (string * string list) list
fe7640933a47 match_bind: 'as' patterns;
wenzelm
parents: 5925
diff changeset
    52
    -> ProofHistory.T -> ProofHistory.T
6371
8469852acbc0 added '_i' versions;
wenzelm
parents: 6354
diff changeset
    53
  val assume_i: string -> Proof.context attribute list -> (term * term list) list
8469852acbc0 added '_i' versions;
wenzelm
parents: 6354
diff changeset
    54
    -> ProofHistory.T -> ProofHistory.T
5938
fe7640933a47 match_bind: 'as' patterns;
wenzelm
parents: 5925
diff changeset
    55
  val show: string -> Args.src list -> string * string list -> ProofHistory.T -> ProofHistory.T
6371
8469852acbc0 added '_i' versions;
wenzelm
parents: 6354
diff changeset
    56
  val show_i: string -> Proof.context attribute list -> term * term list
8469852acbc0 added '_i' versions;
wenzelm
parents: 6354
diff changeset
    57
    -> ProofHistory.T -> ProofHistory.T
5938
fe7640933a47 match_bind: 'as' patterns;
wenzelm
parents: 5925
diff changeset
    58
  val have: string -> Args.src list -> string * string list -> ProofHistory.T -> ProofHistory.T
6371
8469852acbc0 added '_i' versions;
wenzelm
parents: 6354
diff changeset
    59
  val have_i: string -> Proof.context attribute list -> term * term list
8469852acbc0 added '_i' versions;
wenzelm
parents: 6354
diff changeset
    60
    -> ProofHistory.T -> ProofHistory.T
6501
392333eb31cb added thus, hence;
wenzelm
parents: 6483
diff changeset
    61
  val thus: string -> Args.src list -> string * string list -> ProofHistory.T -> ProofHistory.T
392333eb31cb added thus, hence;
wenzelm
parents: 6483
diff changeset
    62
  val thus_i: string -> Proof.context attribute list -> term * term list
392333eb31cb added thus, hence;
wenzelm
parents: 6483
diff changeset
    63
    -> ProofHistory.T -> ProofHistory.T
392333eb31cb added thus, hence;
wenzelm
parents: 6483
diff changeset
    64
  val hence: string -> Args.src list -> string * string list -> ProofHistory.T -> ProofHistory.T
392333eb31cb added thus, hence;
wenzelm
parents: 6483
diff changeset
    65
  val hence_i: string -> Proof.context attribute list -> term * term list
392333eb31cb added thus, hence;
wenzelm
parents: 6483
diff changeset
    66
    -> ProofHistory.T -> ProofHistory.T
5830
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
    67
  val begin_block: ProofHistory.T -> ProofHistory.T
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
    68
  val next_block: ProofHistory.T -> ProofHistory.T
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
    69
  val end_block: ProofHistory.T -> ProofHistory.T
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
    70
  val tac: Method.text -> ProofHistory.T -> ProofHistory.T
5882
c8096461f5c8 renamed tac / etac to refine / then_refine;
wenzelm
parents: 5830
diff changeset
    71
  val then_tac: Method.text -> ProofHistory.T -> ProofHistory.T
5830
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
    72
  val proof: Method.text option -> ProofHistory.T -> ProofHistory.T
6404
2daaf2943c79 common qed and end of proofs;
wenzelm
parents: 6371
diff changeset
    73
  val kill_proof: ProofHistory.T -> theory
2daaf2943c79 common qed and end of proofs;
wenzelm
parents: 6371
diff changeset
    74
  val global_qed_with: bstring option * Args.src list option -> Method.text option
2daaf2943c79 common qed and end of proofs;
wenzelm
parents: 6371
diff changeset
    75
    -> Toplevel.transition -> Toplevel.transition
2daaf2943c79 common qed and end of proofs;
wenzelm
parents: 6371
diff changeset
    76
  val global_qed_with_i: bstring option * theory attribute list option -> Method.text option
2daaf2943c79 common qed and end of proofs;
wenzelm
parents: 6371
diff changeset
    77
    -> Toplevel.transition -> Toplevel.transition
2daaf2943c79 common qed and end of proofs;
wenzelm
parents: 6371
diff changeset
    78
  val qed: Method.text option -> Toplevel.transition -> Toplevel.transition
2daaf2943c79 common qed and end of proofs;
wenzelm
parents: 6371
diff changeset
    79
  val terminal_proof: Method.text -> Toplevel.transition -> Toplevel.transition
2daaf2943c79 common qed and end of proofs;
wenzelm
parents: 6371
diff changeset
    80
  val immediate_proof: Toplevel.transition -> Toplevel.transition
2daaf2943c79 common qed and end of proofs;
wenzelm
parents: 6371
diff changeset
    81
  val default_proof: Toplevel.transition -> Toplevel.transition
5830
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
    82
  val use_mltext: string -> theory option -> theory option
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
    83
  val use_mltext_theory: string -> theory -> theory
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
    84
  val use_setup: string -> theory -> theory
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
    85
  val parse_ast_translation: string -> theory -> theory
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
    86
  val parse_translation: string -> theory -> theory
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
    87
  val print_translation: string -> theory -> theory
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
    88
  val typed_print_translation: string -> theory -> theory
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
    89
  val print_ast_translation: string -> theory -> theory
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
    90
  val token_translation: string -> theory -> theory
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
    91
  val add_oracle: bstring * string -> theory -> theory
6331
fb7b8d6c2bd1 begin/end_theory: presentation;
wenzelm
parents: 6266
diff changeset
    92
  val begin_theory: string -> string list -> (string * bool) list -> theory
fb7b8d6c2bd1 begin/end_theory: presentation;
wenzelm
parents: 6266
diff changeset
    93
  val end_theory: theory -> theory
6246
0aa2e536bc20 improved theory, context, update_context;
wenzelm
parents: 6198
diff changeset
    94
  val theory: string * string list * (string * bool) list
0aa2e536bc20 improved theory, context, update_context;
wenzelm
parents: 6198
diff changeset
    95
    -> Toplevel.transition -> Toplevel.transition
0aa2e536bc20 improved theory, context, update_context;
wenzelm
parents: 6198
diff changeset
    96
  val context: string -> Toplevel.transition -> Toplevel.transition
0aa2e536bc20 improved theory, context, update_context;
wenzelm
parents: 6198
diff changeset
    97
  val update_context: string -> Toplevel.transition -> Toplevel.transition
5830
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
    98
end;
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
    99
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   100
structure IsarThy: ISAR_THY =
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   101
struct
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   102
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   103
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   104
(** derived theory and proof operations **)
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   105
6371
8469852acbc0 added '_i' versions;
wenzelm
parents: 6354
diff changeset
   106
(* formal comments *)   (* FIXME dummy *)
5959
7af99477751b add_text, add_chapter etc.: dummy;
wenzelm
parents: 5938
diff changeset
   107
7af99477751b add_text, add_chapter etc.: dummy;
wenzelm
parents: 5938
diff changeset
   108
fun add_text (txt:string) (thy:theory) = thy;
6354
a4c75cbd2fbf add_title;
wenzelm
parents: 6331
diff changeset
   109
fun add_title title author date thy = thy;
5959
7af99477751b add_text, add_chapter etc.: dummy;
wenzelm
parents: 5938
diff changeset
   110
val add_chapter = add_text;
7af99477751b add_text, add_chapter etc.: dummy;
wenzelm
parents: 5938
diff changeset
   111
val add_section = add_text;
7af99477751b add_text, add_chapter etc.: dummy;
wenzelm
parents: 5938
diff changeset
   112
val add_subsection = add_text;
7af99477751b add_text, add_chapter etc.: dummy;
wenzelm
parents: 5938
diff changeset
   113
val add_subsubsection = add_text;
7af99477751b add_text, add_chapter etc.: dummy;
wenzelm
parents: 5938
diff changeset
   114
7af99477751b add_text, add_chapter etc.: dummy;
wenzelm
parents: 5938
diff changeset
   115
5830
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   116
(* axioms and defs *)
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   117
5915
66f4bde4c6e7 added have_theorems, have_lemmas, have_facts;
wenzelm
parents: 5882
diff changeset
   118
fun add_axms f args thy =
66f4bde4c6e7 added have_theorems, have_lemmas, have_facts;
wenzelm
parents: 5882
diff changeset
   119
  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
   120
66f4bde4c6e7 added have_theorems, have_lemmas, have_facts;
wenzelm
parents: 5882
diff changeset
   121
val add_axioms = add_axms PureThy.add_axioms;
6371
8469852acbc0 added '_i' versions;
wenzelm
parents: 6354
diff changeset
   122
val add_axioms_i = PureThy.add_axioms_i;
5915
66f4bde4c6e7 added have_theorems, have_lemmas, have_facts;
wenzelm
parents: 5882
diff changeset
   123
val add_defs = add_axms PureThy.add_defs;
6371
8469852acbc0 added '_i' versions;
wenzelm
parents: 6354
diff changeset
   124
val add_defs_i = PureThy.add_defs_i;
8469852acbc0 added '_i' versions;
wenzelm
parents: 6354
diff changeset
   125
8469852acbc0 added '_i' versions;
wenzelm
parents: 6354
diff changeset
   126
8469852acbc0 added '_i' versions;
wenzelm
parents: 6354
diff changeset
   127
(* constdefs *)
8469852acbc0 added '_i' versions;
wenzelm
parents: 6354
diff changeset
   128
8469852acbc0 added '_i' versions;
wenzelm
parents: 6354
diff changeset
   129
fun gen_add_constdefs consts defs args thy =
8469852acbc0 added '_i' versions;
wenzelm
parents: 6354
diff changeset
   130
  thy
8469852acbc0 added '_i' versions;
wenzelm
parents: 6354
diff changeset
   131
  |> consts (map fst args)
8469852acbc0 added '_i' versions;
wenzelm
parents: 6354
diff changeset
   132
  |> defs (map (fn ((c, _, mx), s) => ((Thm.def_name (Syntax.const_name c mx), s), [])) args);
8469852acbc0 added '_i' versions;
wenzelm
parents: 6354
diff changeset
   133
8469852acbc0 added '_i' versions;
wenzelm
parents: 6354
diff changeset
   134
val add_constdefs = gen_add_constdefs Theory.add_consts add_defs;
8469852acbc0 added '_i' versions;
wenzelm
parents: 6354
diff changeset
   135
val add_constdefs_i = gen_add_constdefs Theory.add_consts_i add_defs_i;
5915
66f4bde4c6e7 added have_theorems, have_lemmas, have_facts;
wenzelm
parents: 5882
diff changeset
   136
66f4bde4c6e7 added have_theorems, have_lemmas, have_facts;
wenzelm
parents: 5882
diff changeset
   137
66f4bde4c6e7 added have_theorems, have_lemmas, have_facts;
wenzelm
parents: 5882
diff changeset
   138
(* theorems *)
66f4bde4c6e7 added have_theorems, have_lemmas, have_facts;
wenzelm
parents: 5882
diff changeset
   139
66f4bde4c6e7 added have_theorems, have_lemmas, have_facts;
wenzelm
parents: 5882
diff changeset
   140
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
   141
  f name (map (attrib x) more_srcs)
66f4bde4c6e7 added have_theorems, have_lemmas, have_facts;
wenzelm
parents: 5882
diff changeset
   142
    (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
   143
6371
8469852acbc0 added '_i' versions;
wenzelm
parents: 6354
diff changeset
   144
fun global_have_thmss x = gen_have_thmss PureThy.get_thms Attrib.global_attribute x;
5915
66f4bde4c6e7 added have_theorems, have_lemmas, have_facts;
wenzelm
parents: 5882
diff changeset
   145
6371
8469852acbc0 added '_i' versions;
wenzelm
parents: 6354
diff changeset
   146
fun local_have_thmss x =
8469852acbc0 added '_i' versions;
wenzelm
parents: 6354
diff changeset
   147
  gen_have_thmss (ProofContext.get_thms o Proof.context_of)
8469852acbc0 added '_i' versions;
wenzelm
parents: 6354
diff changeset
   148
    (Attrib.local_attribute o Proof.theory_of) x;
8469852acbc0 added '_i' versions;
wenzelm
parents: 6354
diff changeset
   149
8469852acbc0 added '_i' versions;
wenzelm
parents: 6354
diff changeset
   150
fun have_thmss_i f ((name, more_atts), th_atts) =
8469852acbc0 added '_i' versions;
wenzelm
parents: 6354
diff changeset
   151
  f name more_atts (map (apfst single) th_atts);
8469852acbc0 added '_i' versions;
wenzelm
parents: 6354
diff changeset
   152
8469852acbc0 added '_i' versions;
wenzelm
parents: 6354
diff changeset
   153
fun have_lemss name atts = PureThy.have_thmss name (atts @ [Drule.tag_lemma]);
5915
66f4bde4c6e7 added have_theorems, have_lemmas, have_facts;
wenzelm
parents: 5882
diff changeset
   154
66f4bde4c6e7 added have_theorems, have_lemmas, have_facts;
wenzelm
parents: 5882
diff changeset
   155
6371
8469852acbc0 added '_i' versions;
wenzelm
parents: 6354
diff changeset
   156
fun apply_theorems th_srcs = global_have_thmss PureThy.have_thmss ((None, []), th_srcs);
8469852acbc0 added '_i' versions;
wenzelm
parents: 6354
diff changeset
   157
fun apply_theorems_i th_srcs = have_thmss_i PureThy.have_thmss ((None, []), th_srcs);
8469852acbc0 added '_i' versions;
wenzelm
parents: 6354
diff changeset
   158
val have_theorems = #1 oo global_have_thmss (PureThy.have_thmss o Some);
8469852acbc0 added '_i' versions;
wenzelm
parents: 6354
diff changeset
   159
val have_theorems_i = #1 oo have_thmss_i (PureThy.have_thmss o Some);
8469852acbc0 added '_i' versions;
wenzelm
parents: 6354
diff changeset
   160
val have_lemmas = #1 oo global_have_thmss (have_lemss o Some);
8469852acbc0 added '_i' versions;
wenzelm
parents: 6354
diff changeset
   161
val have_lemmas_i = #1 oo have_thmss_i (have_lemss o Some);
8469852acbc0 added '_i' versions;
wenzelm
parents: 6354
diff changeset
   162
val have_facts = ProofHistory.apply o local_have_thmss Proof.have_thmss;
8469852acbc0 added '_i' versions;
wenzelm
parents: 6354
diff changeset
   163
val have_facts_i = ProofHistory.apply o have_thmss_i Proof.have_thmss;
8469852acbc0 added '_i' versions;
wenzelm
parents: 6354
diff changeset
   164
5915
66f4bde4c6e7 added have_theorems, have_lemmas, have_facts;
wenzelm
parents: 5882
diff changeset
   165
6371
8469852acbc0 added '_i' versions;
wenzelm
parents: 6354
diff changeset
   166
(* forward chaining *)
8469852acbc0 added '_i' versions;
wenzelm
parents: 6354
diff changeset
   167
8469852acbc0 added '_i' versions;
wenzelm
parents: 6354
diff changeset
   168
val from_facts =
8469852acbc0 added '_i' versions;
wenzelm
parents: 6354
diff changeset
   169
  ProofHistory.apply o (Proof.chain oo curry (local_have_thmss Proof.have_thmss) ("", []));
8469852acbc0 added '_i' versions;
wenzelm
parents: 6354
diff changeset
   170
8469852acbc0 added '_i' versions;
wenzelm
parents: 6354
diff changeset
   171
val from_facts_i =
8469852acbc0 added '_i' versions;
wenzelm
parents: 6354
diff changeset
   172
  ProofHistory.apply o (Proof.chain oo curry (have_thmss_i Proof.have_thmss) ("", []));
8469852acbc0 added '_i' versions;
wenzelm
parents: 6354
diff changeset
   173
8469852acbc0 added '_i' versions;
wenzelm
parents: 6354
diff changeset
   174
val chain = ProofHistory.apply Proof.chain;
5830
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   175
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   176
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   177
(* context *)
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   178
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   179
val fix = ProofHistory.apply o Proof.fix;
6371
8469852acbc0 added '_i' versions;
wenzelm
parents: 6354
diff changeset
   180
val fix_i = ProofHistory.apply o Proof.fix_i;
5830
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   181
val match_bind = ProofHistory.apply o Proof.match_bind;
6371
8469852acbc0 added '_i' versions;
wenzelm
parents: 6354
diff changeset
   182
val match_bind_i = ProofHistory.apply o Proof.match_bind_i;
5830
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   183
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   184
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   185
(* statements *)
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   186
6371
8469852acbc0 added '_i' versions;
wenzelm
parents: 6354
diff changeset
   187
fun global_statement f name src s thy =
8469852acbc0 added '_i' versions;
wenzelm
parents: 6354
diff changeset
   188
  ProofHistory.init (f name (map (Attrib.global_attribute thy) src) s thy);
5830
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   189
6501
392333eb31cb added thus, hence;
wenzelm
parents: 6483
diff changeset
   190
fun global_statement_i f name atts t thy = ProofHistory.init (f name atts t thy);
392333eb31cb added thus, hence;
wenzelm
parents: 6483
diff changeset
   191
392333eb31cb added thus, hence;
wenzelm
parents: 6483
diff changeset
   192
fun local_statement do_open f g name src s = ProofHistory.apply_cond_open do_open (fn state =>
392333eb31cb added thus, hence;
wenzelm
parents: 6483
diff changeset
   193
  f name (map (Attrib.local_attribute (Proof.theory_of state)) src) s (g state));
6371
8469852acbc0 added '_i' versions;
wenzelm
parents: 6354
diff changeset
   194
6501
392333eb31cb added thus, hence;
wenzelm
parents: 6483
diff changeset
   195
fun local_statement_i do_open f g name atts t =
392333eb31cb added thus, hence;
wenzelm
parents: 6483
diff changeset
   196
  ProofHistory.apply_cond_open do_open (f name atts t o g);
5830
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   197
6501
392333eb31cb added thus, hence;
wenzelm
parents: 6483
diff changeset
   198
val theorem   = global_statement Proof.theorem;
6371
8469852acbc0 added '_i' versions;
wenzelm
parents: 6354
diff changeset
   199
val theorem_i = global_statement_i Proof.theorem_i;
6501
392333eb31cb added thus, hence;
wenzelm
parents: 6483
diff changeset
   200
val lemma     = global_statement Proof.lemma;
392333eb31cb added thus, hence;
wenzelm
parents: 6483
diff changeset
   201
val lemma_i   = global_statement_i Proof.lemma_i;
392333eb31cb added thus, hence;
wenzelm
parents: 6483
diff changeset
   202
val assume    = local_statement false Proof.assume I;
392333eb31cb added thus, hence;
wenzelm
parents: 6483
diff changeset
   203
val assume_i  = local_statement_i false Proof.assume_i I;
392333eb31cb added thus, hence;
wenzelm
parents: 6483
diff changeset
   204
val show      = local_statement true Proof.show I;
392333eb31cb added thus, hence;
wenzelm
parents: 6483
diff changeset
   205
val show_i    = local_statement_i true Proof.show_i I;
392333eb31cb added thus, hence;
wenzelm
parents: 6483
diff changeset
   206
val have      = local_statement true Proof.have I;
392333eb31cb added thus, hence;
wenzelm
parents: 6483
diff changeset
   207
val have_i    = local_statement_i true Proof.have_i I;
392333eb31cb added thus, hence;
wenzelm
parents: 6483
diff changeset
   208
val thus      = local_statement true Proof.show Proof.chain;
392333eb31cb added thus, hence;
wenzelm
parents: 6483
diff changeset
   209
val thus_i    = local_statement_i true Proof.show_i Proof.chain;
392333eb31cb added thus, hence;
wenzelm
parents: 6483
diff changeset
   210
val hence     = local_statement true Proof.have Proof.chain;
392333eb31cb added thus, hence;
wenzelm
parents: 6483
diff changeset
   211
val hence_i   = local_statement_i true Proof.have_i Proof.chain;
5830
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   212
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   213
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   214
(* blocks *)
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   215
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   216
val begin_block = ProofHistory.apply_open Proof.begin_block;
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   217
val next_block = ProofHistory.apply Proof.next_block;
6404
2daaf2943c79 common qed and end of proofs;
wenzelm
parents: 6371
diff changeset
   218
val end_block = ProofHistory.apply_close Proof.end_block;
5830
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   219
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   220
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   221
(* backward steps *)
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   222
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   223
val tac = ProofHistory.applys o Method.tac;
5882
c8096461f5c8 renamed tac / etac to refine / then_refine;
wenzelm
parents: 5830
diff changeset
   224
val then_tac = ProofHistory.applys o Method.then_tac;
5830
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   225
val proof = ProofHistory.applys o Method.proof;
6404
2daaf2943c79 common qed and end of proofs;
wenzelm
parents: 6371
diff changeset
   226
2daaf2943c79 common qed and end of proofs;
wenzelm
parents: 6371
diff changeset
   227
2daaf2943c79 common qed and end of proofs;
wenzelm
parents: 6371
diff changeset
   228
(* local endings *)
2daaf2943c79 common qed and end of proofs;
wenzelm
parents: 6371
diff changeset
   229
2daaf2943c79 common qed and end of proofs;
wenzelm
parents: 6371
diff changeset
   230
val local_qed = Toplevel.proof o ProofHistory.applys_close o Method.local_qed;
2daaf2943c79 common qed and end of proofs;
wenzelm
parents: 6371
diff changeset
   231
val local_terminal_proof = Toplevel.proof o ProofHistory.applys_close o Method.local_terminal_proof;
2daaf2943c79 common qed and end of proofs;
wenzelm
parents: 6371
diff changeset
   232
val local_immediate_proof = Toplevel.proof (ProofHistory.applys_close Method.local_immediate_proof);
2daaf2943c79 common qed and end of proofs;
wenzelm
parents: 6371
diff changeset
   233
val local_default_proof = Toplevel.proof (ProofHistory.applys_close Method.local_default_proof);
2daaf2943c79 common qed and end of proofs;
wenzelm
parents: 6371
diff changeset
   234
2daaf2943c79 common qed and end of proofs;
wenzelm
parents: 6371
diff changeset
   235
2daaf2943c79 common qed and end of proofs;
wenzelm
parents: 6371
diff changeset
   236
(* global endings *)
2daaf2943c79 common qed and end of proofs;
wenzelm
parents: 6371
diff changeset
   237
2daaf2943c79 common qed and end of proofs;
wenzelm
parents: 6371
diff changeset
   238
val kill_proof = Proof.theory_of o ProofHistory.current;
2daaf2943c79 common qed and end of proofs;
wenzelm
parents: 6371
diff changeset
   239
2daaf2943c79 common qed and end of proofs;
wenzelm
parents: 6371
diff changeset
   240
fun global_result finish = Toplevel.proof_to_theory (fn prf =>
2daaf2943c79 common qed and end of proofs;
wenzelm
parents: 6371
diff changeset
   241
  let
2daaf2943c79 common qed and end of proofs;
wenzelm
parents: 6371
diff changeset
   242
    val state = ProofHistory.current prf;
2daaf2943c79 common qed and end of proofs;
wenzelm
parents: 6371
diff changeset
   243
    val _ = if Proof.at_bottom state then () else raise Toplevel.UNDEF;
2daaf2943c79 common qed and end of proofs;
wenzelm
parents: 6371
diff changeset
   244
    val (thy, (kind, name, thm)) = finish state;
2daaf2943c79 common qed and end of proofs;
wenzelm
parents: 6371
diff changeset
   245
2daaf2943c79 common qed and end of proofs;
wenzelm
parents: 6371
diff changeset
   246
    val prt_result = Pretty.block
2daaf2943c79 common qed and end of proofs;
wenzelm
parents: 6371
diff changeset
   247
      [Pretty.str (kind ^ " " ^ name ^ ":"), Pretty.fbrk, Display.pretty_thm thm];
2daaf2943c79 common qed and end of proofs;
wenzelm
parents: 6371
diff changeset
   248
  in Pretty.writeln prt_result; thy end);
2daaf2943c79 common qed and end of proofs;
wenzelm
parents: 6371
diff changeset
   249
2daaf2943c79 common qed and end of proofs;
wenzelm
parents: 6371
diff changeset
   250
fun gen_global_qed_with prep_att (alt_name, raw_atts) opt_text state =
2daaf2943c79 common qed and end of proofs;
wenzelm
parents: 6371
diff changeset
   251
  let
2daaf2943c79 common qed and end of proofs;
wenzelm
parents: 6371
diff changeset
   252
    val thy = Proof.theory_of state;
2daaf2943c79 common qed and end of proofs;
wenzelm
parents: 6371
diff changeset
   253
    val alt_atts = apsome (map (prep_att thy)) raw_atts;
2daaf2943c79 common qed and end of proofs;
wenzelm
parents: 6371
diff changeset
   254
  in Method.global_qed alt_name alt_atts opt_text state end;
2daaf2943c79 common qed and end of proofs;
wenzelm
parents: 6371
diff changeset
   255
2daaf2943c79 common qed and end of proofs;
wenzelm
parents: 6371
diff changeset
   256
val global_qed_with = global_result oo gen_global_qed_with Attrib.global_attribute;
2daaf2943c79 common qed and end of proofs;
wenzelm
parents: 6371
diff changeset
   257
val global_qed_with_i = global_result oo gen_global_qed_with (K I);
2daaf2943c79 common qed and end of proofs;
wenzelm
parents: 6371
diff changeset
   258
val global_qed = global_qed_with (None, None);
2daaf2943c79 common qed and end of proofs;
wenzelm
parents: 6371
diff changeset
   259
2daaf2943c79 common qed and end of proofs;
wenzelm
parents: 6371
diff changeset
   260
val global_terminal_proof = global_result o Method.global_terminal_proof;
2daaf2943c79 common qed and end of proofs;
wenzelm
parents: 6371
diff changeset
   261
val global_immediate_proof = global_result Method.global_immediate_proof;
2daaf2943c79 common qed and end of proofs;
wenzelm
parents: 6371
diff changeset
   262
val global_default_proof = global_result Method.global_default_proof;
2daaf2943c79 common qed and end of proofs;
wenzelm
parents: 6371
diff changeset
   263
2daaf2943c79 common qed and end of proofs;
wenzelm
parents: 6371
diff changeset
   264
2daaf2943c79 common qed and end of proofs;
wenzelm
parents: 6371
diff changeset
   265
(* common endings *)
2daaf2943c79 common qed and end of proofs;
wenzelm
parents: 6371
diff changeset
   266
2daaf2943c79 common qed and end of proofs;
wenzelm
parents: 6371
diff changeset
   267
fun qed opt_text = local_qed opt_text o global_qed opt_text;
2daaf2943c79 common qed and end of proofs;
wenzelm
parents: 6371
diff changeset
   268
fun terminal_proof opt_text = local_terminal_proof opt_text o global_terminal_proof opt_text;
2daaf2943c79 common qed and end of proofs;
wenzelm
parents: 6371
diff changeset
   269
val immediate_proof = local_immediate_proof o global_immediate_proof;
2daaf2943c79 common qed and end of proofs;
wenzelm
parents: 6371
diff changeset
   270
val default_proof = local_default_proof o global_default_proof;
5830
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   271
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   272
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   273
(* use ML text *)
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   274
6331
fb7b8d6c2bd1 begin/end_theory: presentation;
wenzelm
parents: 6266
diff changeset
   275
fun use_mltext txt opt_thy = #2 (Context.pass opt_thy (use_text false) txt);
fb7b8d6c2bd1 begin/end_theory: presentation;
wenzelm
parents: 6266
diff changeset
   276
fun use_mltext_theory txt thy = #2 (Context.pass_theory thy (use_text false) txt);
5830
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   277
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   278
fun use_context txt = use_mltext_theory ("Context.>> (" ^ txt ^ ");");
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   279
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   280
fun use_let name body txt =
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   281
  use_context ("let val " ^ name ^ " = " ^ txt ^ " in\n" ^ body ^ " end");
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   282
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   283
val use_setup =
5915
66f4bde4c6e7 added have_theorems, have_lemmas, have_facts;
wenzelm
parents: 5882
diff changeset
   284
  use_let "setup: (theory -> theory) list" "Library.apply setup";
5830
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   285
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   286
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   287
(* translation functions *)
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   288
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   289
val parse_ast_translation =
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   290
  use_let "parse_ast_translation: (string * (Syntax.ast list -> Syntax.ast)) list"
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   291
    "Theory.add_trfuns (parse_ast_translation, [], [], [])";
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   292
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   293
val parse_translation =
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   294
  use_let "parse_translation: (string * (term list -> term)) list"
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   295
    "Theory.add_trfuns ([], parse_translation, [], [])";
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   296
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   297
val print_translation =
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   298
  use_let "print_translation: (string * (term list -> term)) list"
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   299
    "Theory.add_trfuns ([], [], print_translation, [])";
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   300
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   301
val print_ast_translation =
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   302
  use_let "print_ast_translation: (string * (Syntax.ast list -> Syntax.ast)) list"
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   303
    "Theory.add_trfuns ([], [], [], print_ast_translation)";
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   304
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   305
val typed_print_translation =
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   306
  use_let "typed_print_translation: (string * (bool -> typ -> term list -> term)) list"
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   307
    "Theory.add_trfunsT typed_print_translation";
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   308
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   309
val token_translation =
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   310
  use_let "token_translation: (string * string * (string -> string * int)) list"
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   311
    "Theory.add_tokentrfuns token_translation";
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   312
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   313
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   314
(* add_oracle *)
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   315
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   316
fun add_oracle (name, txt) =
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   317
  use_let
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   318
    "oracle: bstring * (Sign.sg * Object.T -> term)"
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   319
    "Theory.add_oracle oracle"
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   320
    ("(" ^ quote name ^ ", " ^ txt ^ ")");
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   321
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   322
6371
8469852acbc0 added '_i' versions;
wenzelm
parents: 6354
diff changeset
   323
(* theory init and exit *)      (* FIXME move? rearrange? *)
5830
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   324
6331
fb7b8d6c2bd1 begin/end_theory: presentation;
wenzelm
parents: 6266
diff changeset
   325
fun begin_theory name parents files =
6266
a5f9fa6b6d7c Context.fetch, Context.setmp;
wenzelm
parents: 6253
diff changeset
   326
  let
6331
fb7b8d6c2bd1 begin/end_theory: presentation;
wenzelm
parents: 6266
diff changeset
   327
    val paths = map (apfst Path.unpack) files;
fb7b8d6c2bd1 begin/end_theory: presentation;
wenzelm
parents: 6266
diff changeset
   328
    val thy = ThyInfo.begin_theory name parents paths;
fb7b8d6c2bd1 begin/end_theory: presentation;
wenzelm
parents: 6266
diff changeset
   329
  in Present.begin_theory name parents paths; thy end;
5830
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   330
6331
fb7b8d6c2bd1 begin/end_theory: presentation;
wenzelm
parents: 6266
diff changeset
   331
fb7b8d6c2bd1 begin/end_theory: presentation;
wenzelm
parents: 6266
diff changeset
   332
(* FIXME
5830
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   333
fun end_theory thy =
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   334
  let val thy' = PureThy.end_theory thy in
6331
fb7b8d6c2bd1 begin/end_theory: presentation;
wenzelm
parents: 6266
diff changeset
   335
    Present.end_theory (PureThy.get_name thy');
6198
7fa2deb92b39 ThyInfo.begin_theory;
wenzelm
parents: 6108
diff changeset
   336
    transform_error ThyInfo.put_theory thy'
5830
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   337
      handle exn => raise PureThy.ROLLBACK (thy', Some exn)     (* FIXME !!?? *)
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   338
  end;
6331
fb7b8d6c2bd1 begin/end_theory: presentation;
wenzelm
parents: 6266
diff changeset
   339
*)
5830
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   340
6331
fb7b8d6c2bd1 begin/end_theory: presentation;
wenzelm
parents: 6266
diff changeset
   341
fun end_theory thy =
fb7b8d6c2bd1 begin/end_theory: presentation;
wenzelm
parents: 6266
diff changeset
   342
  (Present.end_theory (PureThy.get_name thy); ThyInfo.end_theory thy);
fb7b8d6c2bd1 begin/end_theory: presentation;
wenzelm
parents: 6266
diff changeset
   343
fb7b8d6c2bd1 begin/end_theory: presentation;
wenzelm
parents: 6266
diff changeset
   344
fun bg_theory (name, parents, files) () = begin_theory name parents files;
fb7b8d6c2bd1 begin/end_theory: presentation;
wenzelm
parents: 6266
diff changeset
   345
fun en_theory thy = (end_theory thy; ());
fb7b8d6c2bd1 begin/end_theory: presentation;
wenzelm
parents: 6266
diff changeset
   346
fb7b8d6c2bd1 begin/end_theory: presentation;
wenzelm
parents: 6266
diff changeset
   347
fun theory spec = Toplevel.init_theory (bg_theory spec) en_theory;
6246
0aa2e536bc20 improved theory, context, update_context;
wenzelm
parents: 6198
diff changeset
   348
0aa2e536bc20 improved theory, context, update_context;
wenzelm
parents: 6198
diff changeset
   349
0aa2e536bc20 improved theory, context, update_context;
wenzelm
parents: 6198
diff changeset
   350
(* context switch *)
0aa2e536bc20 improved theory, context, update_context;
wenzelm
parents: 6198
diff changeset
   351
6483
3e5d450c2b31 switch_theory: Context.pass;
wenzelm
parents: 6404
diff changeset
   352
fun switch_theory load s =
3e5d450c2b31 switch_theory: Context.pass;
wenzelm
parents: 6404
diff changeset
   353
  Toplevel.init_theory (fn () => (the (#2 (Context.pass None load s)))) (K ());
6246
0aa2e536bc20 improved theory, context, update_context;
wenzelm
parents: 6198
diff changeset
   354
0aa2e536bc20 improved theory, context, update_context;
wenzelm
parents: 6198
diff changeset
   355
val context = switch_theory ThyInfo.use_thy;
0aa2e536bc20 improved theory, context, update_context;
wenzelm
parents: 6198
diff changeset
   356
val update_context = switch_theory ThyInfo.update_thy;
0aa2e536bc20 improved theory, context, update_context;
wenzelm
parents: 6198
diff changeset
   357
5830
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   358
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   359
end;