author | berghofe |
Mon, 31 Dec 2001 14:08:23 +0100 | |
changeset 12611 | c44a9fecb518 |
parent 11895 | 73b2c277974f |
child 13279 | 8a722689a1c9 |
permissions | -rw-r--r-- |
1132
dfb29abcf3c2
added theorem database which contains axioms and theorems indexed by the
clasohm
parents:
diff
changeset
|
1 |
(* Title: Pure/Thy/thm_database.ML |
dfb29abcf3c2
added theorem database which contains axioms and theorems indexed by the
clasohm
parents:
diff
changeset
|
2 |
ID: $Id$ |
11895 | 3 |
Author: Markus Wenzel, TU Muenchen |
4 |
License: GPL (GNU GENERAL PUBLIC LICENSE) |
|
1132
dfb29abcf3c2
added theorem database which contains axioms and theorems indexed by the
clasohm
parents:
diff
changeset
|
5 |
|
11895 | 6 |
Interface to thm database (see also Pure/pure_thy.ML). |
4023 | 7 |
*) |
1262
8f40ff1299d8
added removal of theorems if theory is to be reloaded; changed functions for
clasohm
parents:
1245
diff
changeset
|
8 |
|
6204 | 9 |
signature BASIC_THM_DATABASE = |
4023 | 10 |
sig |
11 |
val store_thm: string * thm -> thm |
|
7410 | 12 |
val store_thms: string * thm list -> thm list |
4023 | 13 |
end; |
1132
dfb29abcf3c2
added theorem database which contains axioms and theorems indexed by the
clasohm
parents:
diff
changeset
|
14 |
|
6204 | 15 |
signature THM_DATABASE = |
16 |
sig |
|
17 |
include BASIC_THM_DATABASE |
|
7410 | 18 |
val qed_thms: thm list ref |
6204 | 19 |
val ml_store_thm: string * thm -> unit |
7410 | 20 |
val ml_store_thms: string * thm list -> unit |
6204 | 21 |
val is_ml_identifier: string -> bool |
11529 | 22 |
val ml_reserved: string list |
11895 | 23 |
val thms_containing: theory -> string list -> (string * thm) list |
11017 | 24 |
val print_thms_containing: theory -> term list -> unit |
6204 | 25 |
end; |
26 |
||
3627 | 27 |
structure ThmDatabase: THM_DATABASE = |
1221 | 28 |
struct |
1132
dfb29abcf3c2
added theorem database which contains axioms and theorems indexed by the
clasohm
parents:
diff
changeset
|
29 |
|
4023 | 30 |
(** store theorems **) |
1132
dfb29abcf3c2
added theorem database which contains axioms and theorems indexed by the
clasohm
parents:
diff
changeset
|
31 |
|
4023 | 32 |
(* store in theory and generate HTML *) |
1136
3910c96551d1
fixed bug in thms_containing; changed error/warning messages;
clasohm
parents:
1134
diff
changeset
|
33 |
|
4023 | 34 |
fun store_thm (name, thm) = |
7410 | 35 |
let val thm' = hd (PureThy.smart_store_thms (name, [thm])) |
6327 | 36 |
in Present.theorem name thm'; thm' end; |
1132
dfb29abcf3c2
added theorem database which contains axioms and theorems indexed by the
clasohm
parents:
diff
changeset
|
37 |
|
7410 | 38 |
fun store_thms (name, thms) = |
39 |
let val thms' = PureThy.smart_store_thms (name, thms) |
|
40 |
in Present.theorems name thms'; thms' end; |
|
41 |
||
1262
8f40ff1299d8
added removal of theorems if theory is to be reloaded; changed functions for
clasohm
parents:
1245
diff
changeset
|
42 |
|
4023 | 43 |
(* store on ML toplevel *) |
44 |
||
7410 | 45 |
val qed_thms: thm list ref = ref []; |
1262
8f40ff1299d8
added removal of theorems if theory is to be reloaded; changed functions for
clasohm
parents:
1245
diff
changeset
|
46 |
|
4023 | 47 |
val ml_reserved = |
48 |
["abstype", "and", "andalso", "as", "case", "do", "datatype", "else", |
|
49 |
"end", "exception", "fn", "fun", "handle", "if", "in", "infix", |
|
50 |
"infixr", "let", "local", "nonfix", "of", "op", "open", "orelse", |
|
51 |
"raise", "rec", "then", "type", "val", "with", "withtype", "while", |
|
52 |
"eqtype", "functor", "include", "sharing", "sig", "signature", |
|
53 |
"struct", "structure", "where"]; |
|
1262
8f40ff1299d8
added removal of theorems if theory is to be reloaded; changed functions for
clasohm
parents:
1245
diff
changeset
|
54 |
|
4023 | 55 |
fun is_ml_identifier name = |
56 |
Syntax.is_identifier name andalso not (name mem ml_reserved); |
|
1262
8f40ff1299d8
added removal of theorems if theory is to be reloaded; changed functions for
clasohm
parents:
1245
diff
changeset
|
57 |
|
7410 | 58 |
fun warn_ml name = |
59 |
if is_ml_identifier name then false |
|
7573 | 60 |
else if name = "" then true |
7410 | 61 |
else (warning ("Cannot bind theorem(s) " ^ quote name ^ " as ML value"); true); |
62 |
||
10914 | 63 |
val use_text_verbose = use_text Context.ml_output true; |
7854 | 64 |
|
4023 | 65 |
fun ml_store_thm (name, thm) = |
66 |
let val thm' = store_thm (name, thm) in |
|
7410 | 67 |
if warn_ml name then () |
7854 | 68 |
else (qed_thms := [thm']; use_text_verbose ("val " ^ name ^ " = hd (! ThmDatabase.qed_thms);")) |
7410 | 69 |
end; |
70 |
||
71 |
fun ml_store_thms (name, thms) = |
|
72 |
let val thms' = store_thms (name, thms) in |
|
73 |
if warn_ml name then () |
|
7854 | 74 |
else (qed_thms := thms'; use_text_verbose ("val " ^ name ^ " = ! ThmDatabase.qed_thms;")) |
1262
8f40ff1299d8
added removal of theorems if theory is to be reloaded; changed functions for
clasohm
parents:
1245
diff
changeset
|
75 |
end; |
8f40ff1299d8
added removal of theorems if theory is to be reloaded; changed functions for
clasohm
parents:
1245
diff
changeset
|
76 |
|
4023 | 77 |
|
78 |
||
4037 | 79 |
(** retrieve theorems **) |
1132
dfb29abcf3c2
added theorem database which contains axioms and theorems indexed by the
clasohm
parents:
diff
changeset
|
80 |
|
11895 | 81 |
fun thms_containing thy consts = PureThy.thms_containing thy consts |
11017 | 82 |
handle THEORY (msg, _) => error msg | THM (msg, _, _) => error msg; |
83 |
||
84 |
fun print_thms_containing thy ts = |
|
85 |
let |
|
86 |
val prt_const = |
|
87 |
Pretty.quote o Pretty.str o Sign.cond_extern (Theory.sign_of thy) Sign.constK; |
|
88 |
fun prt_thm (a, th) = Pretty.block [Pretty.str (a ^ ":"), |
|
89 |
Pretty.brk 1, Display.pretty_thm_no_quote (#1 (Drule.freeze_thaw th))]; |
|
7609 | 90 |
|
11017 | 91 |
val cs = foldr Term.add_term_consts (ts, []); |
92 |
val header = |
|
93 |
if Library.null cs then [] |
|
94 |
else [Pretty.block (Pretty.breaks (Pretty.str "theorems containing constants:" :: |
|
95 |
map prt_const cs)), Pretty.str ""]; |
|
96 |
in |
|
97 |
if Library.null cs andalso not (Library.null ts) then |
|
98 |
warning "thms_containing: no constants in specification" |
|
11895 | 99 |
else (header @ map prt_thm (thms_containing thy cs)) |> Pretty.chunks |> Pretty.writeln |
11017 | 100 |
end; |
1132
dfb29abcf3c2
added theorem database which contains axioms and theorems indexed by the
clasohm
parents:
diff
changeset
|
101 |
|
3601
43c7912aac8d
Moved "qed" - functions etc. from thy_read.ML to thm_database.ML.
berghofe
parents:
2150
diff
changeset
|
102 |
|
1132
dfb29abcf3c2
added theorem database which contains axioms and theorems indexed by the
clasohm
parents:
diff
changeset
|
103 |
end; |
6204 | 104 |
|
105 |
structure BasicThmDatabase: BASIC_THM_DATABASE = ThmDatabase; |
|
106 |
open BasicThmDatabase; |