src/Pure/Isar/isar_thy.ML
author wenzelm
Thu, 27 Jul 2006 13:43:01 +0200
changeset 20224 9c40a144ee0e
parent 19632 21e04f0edd82
child 20363 f34c5dbe74d5
permissions -rw-r--r--
moved basic assumption operations from structure ProofContext to Assumption;
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
17354
4d92517aa7f4 moved most material to its proper place (sign.ML, pure_thy.ML, method.ML, proof.ML, locale.ML etc.);
wenzelm
parents: 17184
diff changeset
     5
Derived theory and proof operations.
5830
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
     6
*)
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
     7
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
     8
signature ISAR_THY =
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
     9
sig
15703
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15696
diff changeset
    10
  val add_axioms: ((bstring * string) * Attrib.src list) list -> theory -> theory
19632
21e04f0edd82 add_defs: unchecked flag;
wenzelm
parents: 19585
diff changeset
    11
  val add_defs: (bool * bool) * ((bstring * string) * Attrib.src list) list -> theory -> theory
18418
bf448d999b7e re-arranged tuples (theory * 'a) to ('a * theory) in Pure
haftmann
parents: 18377
diff changeset
    12
  val apply_theorems: (thmref * Attrib.src list) list -> theory -> thm list * theory
18728
6790126ab5f6 simplified type attribute;
wenzelm
parents: 18678
diff changeset
    13
  val apply_theorems_i: (thm list * attribute list) list -> theory -> thm list * theory
15703
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15696
diff changeset
    14
  val theorems: string ->
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15696
diff changeset
    15
    ((bstring * Attrib.src list) * (thmref * Attrib.src list) list) list
19632
21e04f0edd82 add_defs: unchecked flag;
wenzelm
parents: 19585
diff changeset
    16
    -> theory -> (string * thm list) list * theory
12713
c9e3e34dbc45 clarified/added theorems(_i) vs. locale_theorems(_i);
wenzelm
parents: 12705
diff changeset
    17
  val theorems_i: string ->
18728
6790126ab5f6 simplified type attribute;
wenzelm
parents: 18678
diff changeset
    18
    ((bstring * attribute list) * (thm list * attribute list) list) list
18418
bf448d999b7e re-arranged tuples (theory * 'a) to ('a * theory) in Pure
haftmann
parents: 18377
diff changeset
    19
    -> theory -> (string * thm list) list * theory
12713
c9e3e34dbc45 clarified/added theorems(_i) vs. locale_theorems(_i);
wenzelm
parents: 12705
diff changeset
    20
  val smart_theorems: string -> xstring option ->
17354
4d92517aa7f4 moved most material to its proper place (sign.ML, pure_thy.ML, method.ML, proof.ML, locale.ML etc.);
wenzelm
parents: 17184
diff changeset
    21
    ((bstring * Attrib.src list) * (thmref * Attrib.src list) list) list ->
18804
d42708f5c186 swapped Toplevel.theory_context;
wenzelm
parents: 18728
diff changeset
    22
    theory -> Proof.context * theory
15703
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15696
diff changeset
    23
  val declare_theorems: xstring option ->
18804
d42708f5c186 swapped Toplevel.theory_context;
wenzelm
parents: 18728
diff changeset
    24
    (thmref * Attrib.src list) list -> theory -> Proof.context * theory
19585
70a1ce3b23ae removed 'concl is' patterns;
wenzelm
parents: 19482
diff changeset
    25
  val have: ((string * Attrib.src list) * (string * string list) list) list ->
17354
4d92517aa7f4 moved most material to its proper place (sign.ML, pure_thy.ML, method.ML, proof.ML, locale.ML etc.);
wenzelm
parents: 17184
diff changeset
    26
    bool -> Proof.state -> Proof.state
19585
70a1ce3b23ae removed 'concl is' patterns;
wenzelm
parents: 19482
diff changeset
    27
  val hence: ((string * Attrib.src list) * (string * string list) list) list ->
17354
4d92517aa7f4 moved most material to its proper place (sign.ML, pure_thy.ML, method.ML, proof.ML, locale.ML etc.);
wenzelm
parents: 17184
diff changeset
    28
    bool -> Proof.state -> Proof.state
19585
70a1ce3b23ae removed 'concl is' patterns;
wenzelm
parents: 19482
diff changeset
    29
  val show: ((string * Attrib.src list) * (string * string list) list) list ->
17354
4d92517aa7f4 moved most material to its proper place (sign.ML, pure_thy.ML, method.ML, proof.ML, locale.ML etc.);
wenzelm
parents: 17184
diff changeset
    30
    bool -> Proof.state -> Proof.state
19585
70a1ce3b23ae removed 'concl is' patterns;
wenzelm
parents: 19482
diff changeset
    31
  val thus: ((string * Attrib.src list) * (string * string list) list) list ->
70a1ce3b23ae removed 'concl is' patterns;
wenzelm
parents: 19482
diff changeset
    32
    bool -> Proof.state -> Proof.state
70a1ce3b23ae removed 'concl is' patterns;
wenzelm
parents: 19482
diff changeset
    33
  val theorem: string -> string * Attrib.src list -> string * string list -> theory -> Proof.state
70a1ce3b23ae removed 'concl is' patterns;
wenzelm
parents: 19482
diff changeset
    34
  val theorem_i: string -> string * attribute list -> term * term list -> theory -> Proof.state
12876
a70df1e5bf10 got rid of explicit marginal comments (now stripped earlier from input);
wenzelm
parents: 12848
diff changeset
    35
  val qed: Method.text option -> Toplevel.transition -> Toplevel.transition
17354
4d92517aa7f4 moved most material to its proper place (sign.ML, pure_thy.ML, method.ML, proof.ML, locale.ML etc.);
wenzelm
parents: 17184
diff changeset
    36
  val terminal_proof: Method.text * Method.text option ->
4d92517aa7f4 moved most material to its proper place (sign.ML, pure_thy.ML, method.ML, proof.ML, locale.ML etc.);
wenzelm
parents: 17184
diff changeset
    37
    Toplevel.transition -> Toplevel.transition
12876
a70df1e5bf10 got rid of explicit marginal comments (now stripped earlier from input);
wenzelm
parents: 12848
diff changeset
    38
  val default_proof: Toplevel.transition -> Toplevel.transition
a70df1e5bf10 got rid of explicit marginal comments (now stripped earlier from input);
wenzelm
parents: 12848
diff changeset
    39
  val immediate_proof: Toplevel.transition -> Toplevel.transition
a70df1e5bf10 got rid of explicit marginal comments (now stripped earlier from input);
wenzelm
parents: 12848
diff changeset
    40
  val done_proof: Toplevel.transition -> Toplevel.transition
a70df1e5bf10 got rid of explicit marginal comments (now stripped earlier from input);
wenzelm
parents: 12848
diff changeset
    41
  val skip_proof: Toplevel.transition -> Toplevel.transition
17354
4d92517aa7f4 moved most material to its proper place (sign.ML, pure_thy.ML, method.ML, proof.ML, locale.ML etc.);
wenzelm
parents: 17184
diff changeset
    42
  val begin_theory: string -> string list -> (string * bool) list -> bool -> theory
6331
fb7b8d6c2bd1 begin/end_theory: presentation;
wenzelm
parents: 6266
diff changeset
    43
  val end_theory: theory -> theory
7932
92df50fb89ca export kill_theory;
wenzelm
parents: 7909
diff changeset
    44
  val kill_theory: string -> unit
6246
0aa2e536bc20 improved theory, context, update_context;
wenzelm
parents: 6198
diff changeset
    45
  val theory: string * string list * (string * bool) list
0aa2e536bc20 improved theory, context, update_context;
wenzelm
parents: 6198
diff changeset
    46
    -> Toplevel.transition -> Toplevel.transition
7960
d5c91c131070 improved IsarThy.init_context;
wenzelm
parents: 7953
diff changeset
    47
  val init_context: ('a -> unit) -> 'a -> Toplevel.transition -> Toplevel.transition
6246
0aa2e536bc20 improved theory, context, update_context;
wenzelm
parents: 6198
diff changeset
    48
  val context: string -> Toplevel.transition -> Toplevel.transition
5830
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
    49
end;
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
    50
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
    51
structure IsarThy: ISAR_THY =
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
    52
struct
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
    53
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
    54
(* axioms and defs *)
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
    55
5915
66f4bde4c6e7 added have_theorems, have_lemmas, have_facts;
wenzelm
parents: 5882
diff changeset
    56
fun add_axms f args thy =
18728
6790126ab5f6 simplified type attribute;
wenzelm
parents: 18678
diff changeset
    57
  f (map (fn (x, srcs) => (x, map (Attrib.attribute thy) srcs)) args) thy;
5915
66f4bde4c6e7 added have_theorems, have_lemmas, have_facts;
wenzelm
parents: 5882
diff changeset
    58
18377
0e1d025d57b3 oriented result pairs in PureThy
haftmann
parents: 18358
diff changeset
    59
val add_axioms = add_axms (snd oo PureThy.add_axioms);
19632
21e04f0edd82 add_defs: unchecked flag;
wenzelm
parents: 19585
diff changeset
    60
21e04f0edd82 add_defs: unchecked flag;
wenzelm
parents: 19585
diff changeset
    61
fun add_defs ((unchecked, overloaded), args) =
21e04f0edd82 add_defs: unchecked flag;
wenzelm
parents: 19585
diff changeset
    62
  add_axms
21e04f0edd82 add_defs: unchecked flag;
wenzelm
parents: 19585
diff changeset
    63
    (snd oo (if unchecked then PureThy.add_defs_unchecked else PureThy.add_defs) overloaded) args;
6371
8469852acbc0 added '_i' versions;
wenzelm
parents: 6354
diff changeset
    64
8469852acbc0 added '_i' versions;
wenzelm
parents: 6354
diff changeset
    65
5915
66f4bde4c6e7 added have_theorems, have_lemmas, have_facts;
wenzelm
parents: 5882
diff changeset
    66
(* theorems *)
66f4bde4c6e7 added have_theorems, have_lemmas, have_facts;
wenzelm
parents: 5882
diff changeset
    67
18804
d42708f5c186 swapped Toplevel.theory_context;
wenzelm
parents: 18728
diff changeset
    68
fun present_theorems kind (res, thy) =
d42708f5c186 swapped Toplevel.theory_context;
wenzelm
parents: 18728
diff changeset
    69
  conditional (kind <> "" andalso kind <> PureThy.internalK) (fn () =>
d42708f5c186 swapped Toplevel.theory_context;
wenzelm
parents: 18728
diff changeset
    70
    Context.setmp (SOME thy) (Present.results (kind ^ "s")) res);
17108
3962f74bbb8e moved translation functions to Pure/sign.ML;
wenzelm
parents: 17034
diff changeset
    71
17354
4d92517aa7f4 moved most material to its proper place (sign.ML, pure_thy.ML, method.ML, proof.ML, locale.ML etc.);
wenzelm
parents: 17184
diff changeset
    72
fun theorems kind args thy = thy
18804
d42708f5c186 swapped Toplevel.theory_context;
wenzelm
parents: 18728
diff changeset
    73
  |> PureThy.note_thmss kind (Attrib.map_facts (Attrib.attribute thy) args)
18418
bf448d999b7e re-arranged tuples (theory * 'a) to ('a * theory) in Pure
haftmann
parents: 18377
diff changeset
    74
  |> tap (present_theorems kind);
12713
c9e3e34dbc45 clarified/added theorems(_i) vs. locale_theorems(_i);
wenzelm
parents: 12705
diff changeset
    75
17354
4d92517aa7f4 moved most material to its proper place (sign.ML, pure_thy.ML, method.ML, proof.ML, locale.ML etc.);
wenzelm
parents: 17184
diff changeset
    76
fun theorems_i kind args =
18804
d42708f5c186 swapped Toplevel.theory_context;
wenzelm
parents: 18728
diff changeset
    77
  PureThy.note_thmss_i kind args
18418
bf448d999b7e re-arranged tuples (theory * 'a) to ('a * theory) in Pure
haftmann
parents: 18377
diff changeset
    78
  #> tap (present_theorems kind);
9590
52e71612e42f added declare_theorems(_i);
wenzelm
parents: 9466
diff changeset
    79
19482
9f11af8f7ef9 tuned basic list operators (flat, maps, map_filter);
wenzelm
parents: 18804
diff changeset
    80
fun apply_theorems args = apfst (maps snd) o theorems "" [(("", []), args)];
9f11af8f7ef9 tuned basic list operators (flat, maps, map_filter);
wenzelm
parents: 18804
diff changeset
    81
fun apply_theorems_i args = apfst (maps snd) o theorems_i "" [(("", []), args)];
6371
8469852acbc0 added '_i' versions;
wenzelm
parents: 6354
diff changeset
    82
17354
4d92517aa7f4 moved most material to its proper place (sign.ML, pure_thy.ML, method.ML, proof.ML, locale.ML etc.);
wenzelm
parents: 17184
diff changeset
    83
fun smart_theorems kind NONE args thy = thy
4d92517aa7f4 moved most material to its proper place (sign.ML, pure_thy.ML, method.ML, proof.ML, locale.ML etc.);
wenzelm
parents: 17184
diff changeset
    84
      |> theorems kind args
18418
bf448d999b7e re-arranged tuples (theory * 'a) to ('a * theory) in Pure
haftmann
parents: 18377
diff changeset
    85
      |> tap (present_theorems kind)
18804
d42708f5c186 swapped Toplevel.theory_context;
wenzelm
parents: 18728
diff changeset
    86
      |> (fn (_, thy) => `ProofContext.init thy)
17354
4d92517aa7f4 moved most material to its proper place (sign.ML, pure_thy.ML, method.ML, proof.ML, locale.ML etc.);
wenzelm
parents: 17184
diff changeset
    87
  | smart_theorems kind (SOME loc) args thy = thy
4d92517aa7f4 moved most material to its proper place (sign.ML, pure_thy.ML, method.ML, proof.ML, locale.ML etc.);
wenzelm
parents: 17184
diff changeset
    88
      |> Locale.note_thmss kind loc args
18804
d42708f5c186 swapped Toplevel.theory_context;
wenzelm
parents: 18728
diff changeset
    89
      |> tap (fn ((_, res), (_, thy')) => present_theorems kind (res, thy'))
18421
464c93701351 re-arranged tuples (theory * 'a) to ('a * theory) in Pure
haftmann
parents: 18418
diff changeset
    90
      |> #2;
17354
4d92517aa7f4 moved most material to its proper place (sign.ML, pure_thy.ML, method.ML, proof.ML, locale.ML etc.);
wenzelm
parents: 17184
diff changeset
    91
4d92517aa7f4 moved most material to its proper place (sign.ML, pure_thy.ML, method.ML, proof.ML, locale.ML etc.);
wenzelm
parents: 17184
diff changeset
    92
fun declare_theorems opt_loc args =
4d92517aa7f4 moved most material to its proper place (sign.ML, pure_thy.ML, method.ML, proof.ML, locale.ML etc.);
wenzelm
parents: 17184
diff changeset
    93
  smart_theorems "" opt_loc [(("", []), args)];
12705
wenzelm
parents: 12701
diff changeset
    94
5915
66f4bde4c6e7 added have_theorems, have_lemmas, have_facts;
wenzelm
parents: 5882
diff changeset
    95
17354
4d92517aa7f4 moved most material to its proper place (sign.ML, pure_thy.ML, method.ML, proof.ML, locale.ML etc.);
wenzelm
parents: 17184
diff changeset
    96
(* goals *)
10935
808e9dbc68c4 show/thus: check_goal;
wenzelm
parents: 10464
diff changeset
    97
808e9dbc68c4 show/thus: check_goal;
wenzelm
parents: 10464
diff changeset
    98
local
808e9dbc68c4 show/thus: check_goal;
wenzelm
parents: 10464
diff changeset
    99
17354
4d92517aa7f4 moved most material to its proper place (sign.ML, pure_thy.ML, method.ML, proof.ML, locale.ML etc.);
wenzelm
parents: 17184
diff changeset
   100
fun local_goal opt_chain goal stmt int =
18124
a310c78298f9 simplified after_qed;
wenzelm
parents: 17900
diff changeset
   101
  opt_chain #> goal NONE (K Seq.single) stmt int;
12244
179f142ffb03 improved treatment of common result name;
wenzelm
parents: 12242
diff changeset
   102
17354
4d92517aa7f4 moved most material to its proper place (sign.ML, pure_thy.ML, method.ML, proof.ML, locale.ML etc.);
wenzelm
parents: 17184
diff changeset
   103
fun global_goal goal kind a propp thy =
18124
a310c78298f9 simplified after_qed;
wenzelm
parents: 17900
diff changeset
   104
  goal kind NONE (K I) (SOME "") a [(("", []), [propp])] (ProofContext.init thy);
10935
808e9dbc68c4 show/thus: check_goal;
wenzelm
parents: 10464
diff changeset
   105
808e9dbc68c4 show/thus: check_goal;
wenzelm
parents: 10464
diff changeset
   106
in
808e9dbc68c4 show/thus: check_goal;
wenzelm
parents: 10464
diff changeset
   107
17354
4d92517aa7f4 moved most material to its proper place (sign.ML, pure_thy.ML, method.ML, proof.ML, locale.ML etc.);
wenzelm
parents: 17184
diff changeset
   108
val have = local_goal I Proof.have;
4d92517aa7f4 moved most material to its proper place (sign.ML, pure_thy.ML, method.ML, proof.ML, locale.ML etc.);
wenzelm
parents: 17184
diff changeset
   109
val hence = local_goal Proof.chain Proof.have;
4d92517aa7f4 moved most material to its proper place (sign.ML, pure_thy.ML, method.ML, proof.ML, locale.ML etc.);
wenzelm
parents: 17184
diff changeset
   110
val show = local_goal I Proof.show;
4d92517aa7f4 moved most material to its proper place (sign.ML, pure_thy.ML, method.ML, proof.ML, locale.ML etc.);
wenzelm
parents: 17184
diff changeset
   111
val thus = local_goal Proof.chain Proof.show;
4d92517aa7f4 moved most material to its proper place (sign.ML, pure_thy.ML, method.ML, proof.ML, locale.ML etc.);
wenzelm
parents: 17184
diff changeset
   112
val theorem = global_goal Proof.theorem;
4d92517aa7f4 moved most material to its proper place (sign.ML, pure_thy.ML, method.ML, proof.ML, locale.ML etc.);
wenzelm
parents: 17184
diff changeset
   113
val theorem_i = global_goal Proof.theorem_i;
11890
28e42a90bea8 improved source arrangement of obtain;
wenzelm
parents: 11793
diff changeset
   114
7271
442456b2a8bb assume: multiple args;
wenzelm
parents: 7242
diff changeset
   115
end;
442456b2a8bb assume: multiple args;
wenzelm
parents: 7242
diff changeset
   116
5830
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   117
6404
2daaf2943c79 common qed and end of proofs;
wenzelm
parents: 6371
diff changeset
   118
(* local endings *)
2daaf2943c79 common qed and end of proofs;
wenzelm
parents: 6371
diff changeset
   119
17900
5f44c71c4ca4 use simplified Toplevel.proof etc.;
wenzelm
parents: 17855
diff changeset
   120
fun local_qed m = Toplevel.proofs (Proof.local_qed (m, true));
5f44c71c4ca4 use simplified Toplevel.proof etc.;
wenzelm
parents: 17855
diff changeset
   121
val local_terminal_proof = Toplevel.proofs o Proof.local_terminal_proof;
5f44c71c4ca4 use simplified Toplevel.proof etc.;
wenzelm
parents: 17855
diff changeset
   122
val local_default_proof = Toplevel.proofs Proof.local_default_proof;
5f44c71c4ca4 use simplified Toplevel.proof etc.;
wenzelm
parents: 17855
diff changeset
   123
val local_immediate_proof = Toplevel.proofs Proof.local_immediate_proof;
5f44c71c4ca4 use simplified Toplevel.proof etc.;
wenzelm
parents: 17855
diff changeset
   124
val local_done_proof = Toplevel.proofs Proof.local_done_proof;
5f44c71c4ca4 use simplified Toplevel.proof etc.;
wenzelm
parents: 17855
diff changeset
   125
val local_skip_proof = Toplevel.proofs' Proof.local_skip_proof;
16605
4590c1f79050 added print_mode three_buffersN and corresponding cond_print;
wenzelm
parents: 16529
diff changeset
   126
15237
250e9be7a09d Some changes to allow skipping of proof scripts.
berghofe
parents: 15206
diff changeset
   127
val skip_local_qed =
250e9be7a09d Some changes to allow skipping of proof scripts.
berghofe
parents: 15206
diff changeset
   128
  Toplevel.skip_proof (History.apply (fn i => if i > 1 then i - 1 else raise Toplevel.UNDEF));
6404
2daaf2943c79 common qed and end of proofs;
wenzelm
parents: 6371
diff changeset
   129
2daaf2943c79 common qed and end of proofs;
wenzelm
parents: 6371
diff changeset
   130
2daaf2943c79 common qed and end of proofs;
wenzelm
parents: 6371
diff changeset
   131
(* global endings *)
2daaf2943c79 common qed and end of proofs;
wenzelm
parents: 6371
diff changeset
   132
17900
5f44c71c4ca4 use simplified Toplevel.proof etc.;
wenzelm
parents: 17855
diff changeset
   133
fun global_ending ending = Toplevel.proof_to_theory_context (fn int => fn state =>
18678
dd0c569fa43d sane ERROR handling;
wenzelm
parents: 18562
diff changeset
   134
  if can (Proof.assert_bottom true) state then ending int state
dd0c569fa43d sane ERROR handling;
wenzelm
parents: 18562
diff changeset
   135
  else raise Toplevel.UNDEF);
6404
2daaf2943c79 common qed and end of proofs;
wenzelm
parents: 6371
diff changeset
   136
17354
4d92517aa7f4 moved most material to its proper place (sign.ML, pure_thy.ML, method.ML, proof.ML, locale.ML etc.);
wenzelm
parents: 17184
diff changeset
   137
fun global_qed m = global_ending (K (Proof.global_qed (m, true)));
4d92517aa7f4 moved most material to its proper place (sign.ML, pure_thy.ML, method.ML, proof.ML, locale.ML etc.);
wenzelm
parents: 17184
diff changeset
   138
val global_terminal_proof = global_ending o K o Proof.global_terminal_proof;
4d92517aa7f4 moved most material to its proper place (sign.ML, pure_thy.ML, method.ML, proof.ML, locale.ML etc.);
wenzelm
parents: 17184
diff changeset
   139
val global_default_proof = global_ending (K Proof.global_default_proof);
4d92517aa7f4 moved most material to its proper place (sign.ML, pure_thy.ML, method.ML, proof.ML, locale.ML etc.);
wenzelm
parents: 17184
diff changeset
   140
val global_immediate_proof = global_ending (K Proof.global_immediate_proof);
4d92517aa7f4 moved most material to its proper place (sign.ML, pure_thy.ML, method.ML, proof.ML, locale.ML etc.);
wenzelm
parents: 17184
diff changeset
   141
val global_skip_proof = global_ending Proof.global_skip_proof;
4d92517aa7f4 moved most material to its proper place (sign.ML, pure_thy.ML, method.ML, proof.ML, locale.ML etc.);
wenzelm
parents: 17184
diff changeset
   142
val global_done_proof = global_ending (K Proof.global_done_proof);
4d92517aa7f4 moved most material to its proper place (sign.ML, pure_thy.ML, method.ML, proof.ML, locale.ML etc.);
wenzelm
parents: 17184
diff changeset
   143
17900
5f44c71c4ca4 use simplified Toplevel.proof etc.;
wenzelm
parents: 17855
diff changeset
   144
val skip_global_qed = Toplevel.skip_proof_to_theory (equal 1);
6404
2daaf2943c79 common qed and end of proofs;
wenzelm
parents: 6371
diff changeset
   145
2daaf2943c79 common qed and end of proofs;
wenzelm
parents: 6371
diff changeset
   146
2daaf2943c79 common qed and end of proofs;
wenzelm
parents: 6371
diff changeset
   147
(* common endings *)
2daaf2943c79 common qed and end of proofs;
wenzelm
parents: 6371
diff changeset
   148
17354
4d92517aa7f4 moved most material to its proper place (sign.ML, pure_thy.ML, method.ML, proof.ML, locale.ML etc.);
wenzelm
parents: 17184
diff changeset
   149
fun qed m = local_qed m o global_qed m o skip_local_qed o skip_global_qed;
4d92517aa7f4 moved most material to its proper place (sign.ML, pure_thy.ML, method.ML, proof.ML, locale.ML etc.);
wenzelm
parents: 17184
diff changeset
   150
fun terminal_proof m = local_terminal_proof m o global_terminal_proof m;
12876
a70df1e5bf10 got rid of explicit marginal comments (now stripped earlier from input);
wenzelm
parents: 12848
diff changeset
   151
val default_proof = local_default_proof o global_default_proof;
a70df1e5bf10 got rid of explicit marginal comments (now stripped earlier from input);
wenzelm
parents: 12848
diff changeset
   152
val immediate_proof = local_immediate_proof o global_immediate_proof;
a70df1e5bf10 got rid of explicit marginal comments (now stripped earlier from input);
wenzelm
parents: 12848
diff changeset
   153
val done_proof = local_done_proof o global_done_proof;
a70df1e5bf10 got rid of explicit marginal comments (now stripped earlier from input);
wenzelm
parents: 12848
diff changeset
   154
val skip_proof = local_skip_proof o global_skip_proof;
17354
4d92517aa7f4 moved most material to its proper place (sign.ML, pure_thy.ML, method.ML, proof.ML, locale.ML etc.);
wenzelm
parents: 17184
diff changeset
   155
5830
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   156
6688
f33c89103fb4 cleaned comments;
wenzelm
parents: 6650
diff changeset
   157
(* theory init and exit *)
5830
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   158
17354
4d92517aa7f4 moved most material to its proper place (sign.ML, pure_thy.ML, method.ML, proof.ML, locale.ML etc.);
wenzelm
parents: 17184
diff changeset
   159
fun begin_theory name imports uses =
4d92517aa7f4 moved most material to its proper place (sign.ML, pure_thy.ML, method.ML, proof.ML, locale.ML etc.);
wenzelm
parents: 17184
diff changeset
   160
  ThyInfo.begin_theory Present.begin_theory name imports (map (apfst Path.unpack) uses);
5830
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   161
17354
4d92517aa7f4 moved most material to its proper place (sign.ML, pure_thy.ML, method.ML, proof.ML, locale.ML etc.);
wenzelm
parents: 17184
diff changeset
   162
val end_theory = (ThyInfo.check_known_thy o Context.theory_name) ? ThyInfo.end_theory;
7909
824ea50b8dbb end/kill_theory: check_known_thy;
wenzelm
parents: 7862
diff changeset
   163
7932
92df50fb89ca export kill_theory;
wenzelm
parents: 7909
diff changeset
   164
val kill_theory = ThyInfo.if_known_thy ThyInfo.remove_thy;
6688
f33c89103fb4 cleaned comments;
wenzelm
parents: 6650
diff changeset
   165
17354
4d92517aa7f4 moved most material to its proper place (sign.ML, pure_thy.ML, method.ML, proof.ML, locale.ML etc.);
wenzelm
parents: 17184
diff changeset
   166
fun theory (name, imports, uses) =
4d92517aa7f4 moved most material to its proper place (sign.ML, pure_thy.ML, method.ML, proof.ML, locale.ML etc.);
wenzelm
parents: 17184
diff changeset
   167
  Toplevel.init_theory (begin_theory name imports uses)
4d92517aa7f4 moved most material to its proper place (sign.ML, pure_thy.ML, method.ML, proof.ML, locale.ML etc.);
wenzelm
parents: 17184
diff changeset
   168
    (fn thy => (end_theory thy; ()))
4d92517aa7f4 moved most material to its proper place (sign.ML, pure_thy.ML, method.ML, proof.ML, locale.ML etc.);
wenzelm
parents: 17184
diff changeset
   169
    (kill_theory o Context.theory_name);
6246
0aa2e536bc20 improved theory, context, update_context;
wenzelm
parents: 6198
diff changeset
   170
0aa2e536bc20 improved theory, context, update_context;
wenzelm
parents: 6198
diff changeset
   171
0aa2e536bc20 improved theory, context, update_context;
wenzelm
parents: 6198
diff changeset
   172
(* context switch *)
0aa2e536bc20 improved theory, context, update_context;
wenzelm
parents: 6198
diff changeset
   173
7960
d5c91c131070 improved IsarThy.init_context;
wenzelm
parents: 7953
diff changeset
   174
fun fetch_context f x =
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15456
diff changeset
   175
  (case Context.pass NONE f x of
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15456
diff changeset
   176
    ((), NONE) => raise Toplevel.UNDEF
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15456
diff changeset
   177
  | ((), SOME thy) => thy);
7960
d5c91c131070 improved IsarThy.init_context;
wenzelm
parents: 7953
diff changeset
   178
d5c91c131070 improved IsarThy.init_context;
wenzelm
parents: 7953
diff changeset
   179
fun init_context f x = Toplevel.init_theory (fn _ => fetch_context f x) (K ()) (K ());
7953
955fde69fa7b added init_context;
wenzelm
parents: 7932
diff changeset
   180
955fde69fa7b added init_context;
wenzelm
parents: 7932
diff changeset
   181
val context = init_context (ThyInfo.quiet_update_thy true);
6246
0aa2e536bc20 improved theory, context, update_context;
wenzelm
parents: 6198
diff changeset
   182
5830
95b619c7289b Derived theory operations.
wenzelm
parents:
diff changeset
   183
end;