| author | wenzelm | 
| Tue, 04 May 1999 11:31:29 +0200 | |
| changeset 6572 | e77641d2f4ac | 
| parent 6560 | 1436349f8b28 | 
| child 6574 | a91b4cfd3104 | 
| permissions | -rw-r--r-- | 
| 3987 | 1  | 
(* Title: Pure/pure_thy.ML  | 
2  | 
ID: $Id$  | 
|
3  | 
Author: Markus Wenzel, TU Muenchen  | 
|
4  | 
||
| 5091 | 5  | 
Theorem database, derived theory operations, and the ProtoPure theory.  | 
| 3987 | 6  | 
*)  | 
7  | 
||
| 
4022
 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 
wenzelm 
parents: 
4013 
diff
changeset
 | 
8  | 
signature BASIC_PURE_THY =  | 
| 
 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 
wenzelm 
parents: 
4013 
diff
changeset
 | 
9  | 
sig  | 
| 5000 | 10  | 
val print_theorems: theory -> unit  | 
11  | 
val print_theory: theory -> unit  | 
|
| 
4022
 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 
wenzelm 
parents: 
4013 
diff
changeset
 | 
12  | 
val get_thm: theory -> xstring -> thm  | 
| 
 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 
wenzelm 
parents: 
4013 
diff
changeset
 | 
13  | 
val get_thms: theory -> xstring -> thm list  | 
| 6094 | 14  | 
val get_thmss: theory -> xstring list -> thm list  | 
| 
4022
 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 
wenzelm 
parents: 
4013 
diff
changeset
 | 
15  | 
val thms_of: theory -> (string * thm) list  | 
| 5091 | 16  | 
structure ProtoPure:  | 
17  | 
sig  | 
|
18  | 
val thy: theory  | 
|
19  | 
val flexpair_def: thm  | 
|
20  | 
val Goal_def: thm  | 
|
21  | 
end  | 
|
| 4853 | 22  | 
end;  | 
| 
4022
 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 
wenzelm 
parents: 
4013 
diff
changeset
 | 
23  | 
|
| 3987 | 24  | 
signature PURE_THY =  | 
25  | 
sig  | 
|
| 
4022
 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 
wenzelm 
parents: 
4013 
diff
changeset
 | 
26  | 
include BASIC_PURE_THY  | 
| 6367 | 27  | 
val cond_extern_thm_sg: Sign.sg -> string -> xstring  | 
| 6091 | 28  | 
val thms_closure: theory -> xstring -> thm list option  | 
| 
4022
 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 
wenzelm 
parents: 
4013 
diff
changeset
 | 
29  | 
val thms_containing: theory -> string list -> (string * thm) list  | 
| 5907 | 30  | 
val default_name: string -> string  | 
| 6091 | 31  | 
val store_thm: (bstring * thm) * theory attribute list -> theory -> theory * thm  | 
| 
4022
 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 
wenzelm 
parents: 
4013 
diff
changeset
 | 
32  | 
val smart_store_thm: (bstring * thm) -> thm  | 
| 6091 | 33  | 
val add_thms: ((bstring * thm) * theory attribute list) list -> theory -> theory  | 
34  | 
val add_thmss: ((bstring * thm list) * theory attribute list) list -> theory -> theory  | 
|
| 6367 | 35  | 
val have_thmss: bstring option -> theory attribute list ->  | 
| 6091 | 36  | 
(thm list * theory attribute list) list -> theory -> theory * thm list  | 
| 4853 | 37  | 
val add_axioms: ((bstring * string) * theory attribute list) list -> theory -> theory  | 
38  | 
val add_axioms_i: ((bstring * term) * theory attribute list) list -> theory -> theory  | 
|
39  | 
val add_axiomss: ((bstring * string list) * theory attribute list) list -> theory -> theory  | 
|
40  | 
val add_axiomss_i: ((bstring * term list) * theory attribute list) list -> theory -> theory  | 
|
41  | 
val add_defs: ((bstring * string) * theory attribute list) list -> theory -> theory  | 
|
42  | 
val add_defs_i: ((bstring * term) * theory attribute list) list -> theory -> theory  | 
|
43  | 
val add_defss: ((bstring * string list) * theory attribute list) list -> theory -> theory  | 
|
44  | 
val add_defss_i: ((bstring * term list) * theory attribute list) list -> theory -> theory  | 
|
| 
4963
 
38aa2d56e28c
added get_name, put_name, global_path, local_path, begin_theory,
 
wenzelm 
parents: 
4933 
diff
changeset
 | 
45  | 
val get_name: theory -> string  | 
| 
 
38aa2d56e28c
added get_name, put_name, global_path, local_path, begin_theory,
 
wenzelm 
parents: 
4933 
diff
changeset
 | 
46  | 
val put_name: string -> theory -> theory  | 
| 
 
38aa2d56e28c
added get_name, put_name, global_path, local_path, begin_theory,
 
wenzelm 
parents: 
4933 
diff
changeset
 | 
47  | 
val global_path: theory -> theory  | 
| 
 
38aa2d56e28c
added get_name, put_name, global_path, local_path, begin_theory,
 
wenzelm 
parents: 
4933 
diff
changeset
 | 
48  | 
val local_path: theory -> theory  | 
| 
 
38aa2d56e28c
added get_name, put_name, global_path, local_path, begin_theory,
 
wenzelm 
parents: 
4933 
diff
changeset
 | 
49  | 
val begin_theory: string -> theory list -> theory  | 
| 
 
38aa2d56e28c
added get_name, put_name, global_path, local_path, begin_theory,
 
wenzelm 
parents: 
4933 
diff
changeset
 | 
50  | 
val end_theory: theory -> theory  | 
| 5091 | 51  | 
exception ROLLBACK of theory * exn option  | 
52  | 
val transaction: (theory -> theory) -> theory -> theory  | 
|
| 
4922
 
03b81b6e1baa
added thms_closure: theory -> xstring -> tthm list option;
 
wenzelm 
parents: 
4853 
diff
changeset
 | 
53  | 
val add_typedecls: (bstring * string list * mixfix) list -> theory -> theory  | 
| 
6560
 
1436349f8b28
renamed 'dummy' to 'dummy_pattern' (less dangerous);
 
wenzelm 
parents: 
6547 
diff
changeset
 | 
54  | 
val dummy_patternN: string  | 
| 3987 | 55  | 
end;  | 
56  | 
||
57  | 
structure PureThy: PURE_THY =  | 
|
58  | 
struct  | 
|
59  | 
||
60  | 
||
| 
4922
 
03b81b6e1baa
added thms_closure: theory -> xstring -> tthm list option;
 
wenzelm 
parents: 
4853 
diff
changeset
 | 
61  | 
(*** theorem database ***)  | 
| 3987 | 62  | 
|
| 
4922
 
03b81b6e1baa
added thms_closure: theory -> xstring -> tthm list option;
 
wenzelm 
parents: 
4853 
diff
changeset
 | 
63  | 
(** data kind 'Pure/theorems' **)  | 
| 3987 | 64  | 
|
| 5005 | 65  | 
structure TheoremsDataArgs =  | 
66  | 
struct  | 
|
67  | 
val name = "Pure/theorems";  | 
|
| 3987 | 68  | 
|
| 5005 | 69  | 
type T =  | 
70  | 
    {space: NameSpace.T,
 | 
|
| 6091 | 71  | 
thms_tab: thm list Symtab.table,  | 
72  | 
const_idx: int * (int * thm) list Symtab.table} ref;  | 
|
| 3987 | 73  | 
|
| 4853 | 74  | 
fun mk_empty _ =  | 
| 5005 | 75  | 
    ref {space = NameSpace.empty, thms_tab = Symtab.empty, const_idx = (0, Symtab.empty)} : T;
 | 
| 3987 | 76  | 
|
| 5005 | 77  | 
val empty = mk_empty ();  | 
| 6547 | 78  | 
fun copy (ref x) = ref x;  | 
| 5005 | 79  | 
val prep_ext = mk_empty;  | 
80  | 
val merge = mk_empty;  | 
|
81  | 
||
82  | 
  fun print sg (ref {space, thms_tab, const_idx = _}) =
 | 
|
| 4853 | 83  | 
let  | 
| 6091 | 84  | 
val prt_thm = Display.pretty_thm o Thm.transfer_sg sg;  | 
| 4853 | 85  | 
fun prt_thms (name, [th]) =  | 
86  | 
Pretty.block [Pretty.str (name ^ ":"), Pretty.brk 1, prt_thm th]  | 
|
87  | 
| prt_thms (name, ths) = Pretty.big_list (name ^ ":") (map prt_thm ths);  | 
|
| 3987 | 88  | 
|
| 5175 | 89  | 
val thmss = sort_wrt fst (map (apfst (NameSpace.cond_extern space)) (Symtab.dest thms_tab));  | 
| 4853 | 90  | 
in  | 
91  | 
      Pretty.writeln (Display.pretty_name_space ("theorem name space", space));
 | 
|
92  | 
Pretty.writeln (Pretty.big_list "theorems:" (map prt_thms thmss))  | 
|
93  | 
end;  | 
|
| 3987 | 94  | 
end;  | 
95  | 
||
| 5005 | 96  | 
structure TheoremsData = TheoryDataFun(TheoremsDataArgs);  | 
97  | 
val get_theorems_sg = TheoremsData.get_sg;  | 
|
98  | 
val get_theorems = TheoremsData.get;  | 
|
99  | 
||
| 6367 | 100  | 
val cond_extern_thm_sg = NameSpace.cond_extern o #space o ! o get_theorems_sg;  | 
101  | 
||
| 3987 | 102  | 
|
| 5000 | 103  | 
(* print theory *)  | 
| 3987 | 104  | 
|
| 5005 | 105  | 
val print_theorems = TheoremsData.print;  | 
| 5000 | 106  | 
fun print_theory thy =  | 
107  | 
(Display.print_theory thy; print_theorems thy);  | 
|
| 
4022
 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 
wenzelm 
parents: 
4013 
diff
changeset
 | 
108  | 
|
| 3987 | 109  | 
|
110  | 
||
| 
4022
 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 
wenzelm 
parents: 
4013 
diff
changeset
 | 
111  | 
(** retrieve theorems **)  | 
| 3987 | 112  | 
|
| 
4922
 
03b81b6e1baa
added thms_closure: theory -> xstring -> tthm list option;
 
wenzelm 
parents: 
4853 
diff
changeset
 | 
113  | 
(* thms_closure *)  | 
| 
 
03b81b6e1baa
added thms_closure: theory -> xstring -> tthm list option;
 
wenzelm 
parents: 
4853 
diff
changeset
 | 
114  | 
|
| 
 
03b81b6e1baa
added thms_closure: theory -> xstring -> tthm list option;
 
wenzelm 
parents: 
4853 
diff
changeset
 | 
115  | 
(*note: we avoid life references to the theory, so users may safely  | 
| 5000 | 116  | 
keep thms_closure with moderate space consumption*)  | 
| 
4022
 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 
wenzelm 
parents: 
4013 
diff
changeset
 | 
117  | 
|
| 
4922
 
03b81b6e1baa
added thms_closure: theory -> xstring -> tthm list option;
 
wenzelm 
parents: 
4853 
diff
changeset
 | 
118  | 
fun thms_closure_aux thy =  | 
| 
 
03b81b6e1baa
added thms_closure: theory -> xstring -> tthm list option;
 
wenzelm 
parents: 
4853 
diff
changeset
 | 
119  | 
  let val ref {space, thms_tab, ...} = get_theorems thy
 | 
| 
 
03b81b6e1baa
added thms_closure: theory -> xstring -> tthm list option;
 
wenzelm 
parents: 
4853 
diff
changeset
 | 
120  | 
in fn name => Symtab.lookup (thms_tab, NameSpace.intern space name) end;  | 
| 
 
03b81b6e1baa
added thms_closure: theory -> xstring -> tthm list option;
 
wenzelm 
parents: 
4853 
diff
changeset
 | 
121  | 
|
| 
 
03b81b6e1baa
added thms_closure: theory -> xstring -> tthm list option;
 
wenzelm 
parents: 
4853 
diff
changeset
 | 
122  | 
fun thms_closure thy =  | 
| 
 
03b81b6e1baa
added thms_closure: theory -> xstring -> tthm list option;
 
wenzelm 
parents: 
4853 
diff
changeset
 | 
123  | 
let val closures = map thms_closure_aux (thy :: Theory.ancestors_of thy)  | 
| 
 
03b81b6e1baa
added thms_closure: theory -> xstring -> tthm list option;
 
wenzelm 
parents: 
4853 
diff
changeset
 | 
124  | 
in fn name => get_first (fn f => f name) closures end;  | 
| 
4022
 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 
wenzelm 
parents: 
4013 
diff
changeset
 | 
125  | 
|
| 3987 | 126  | 
|
| 4783 | 127  | 
(* get_thms etc. *)  | 
| 4037 | 128  | 
|
| 
4922
 
03b81b6e1baa
added thms_closure: theory -> xstring -> tthm list option;
 
wenzelm 
parents: 
4853 
diff
changeset
 | 
129  | 
fun lookup_thms name thy = thms_closure_aux thy name;  | 
| 
 
03b81b6e1baa
added thms_closure: theory -> xstring -> tthm list option;
 
wenzelm 
parents: 
4853 
diff
changeset
 | 
130  | 
|
| 6091 | 131  | 
fun get_thms thy name =  | 
| 
4922
 
03b81b6e1baa
added thms_closure: theory -> xstring -> tthm list option;
 
wenzelm 
parents: 
4853 
diff
changeset
 | 
132  | 
(case get_first (lookup_thms name) (thy :: Theory.ancestors_of thy) of  | 
| 4590 | 133  | 
    None => raise THEORY ("Unknown theorem(s) " ^ quote name, [thy])
 | 
| 
4022
 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 
wenzelm 
parents: 
4013 
diff
changeset
 | 
134  | 
| Some thms => thms);  | 
| 3987 | 135  | 
|
| 6091 | 136  | 
fun get_thm thy name =  | 
137  | 
(case get_thms thy name of  | 
|
| 3987 | 138  | 
[thm] => thm  | 
| 4590 | 139  | 
  | _ => raise THEORY ("Single theorem expected " ^ quote name, [thy]));
 | 
| 
4022
 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 
wenzelm 
parents: 
4013 
diff
changeset
 | 
140  | 
|
| 6091 | 141  | 
fun get_thmss thy names = flat (map (get_thms thy) names);  | 
| 4783 | 142  | 
|
| 
4022
 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 
wenzelm 
parents: 
4013 
diff
changeset
 | 
143  | 
|
| 
 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 
wenzelm 
parents: 
4013 
diff
changeset
 | 
144  | 
(* thms_of *)  | 
| 
 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 
wenzelm 
parents: 
4013 
diff
changeset
 | 
145  | 
|
| 6091 | 146  | 
fun attach_name thm = (Thm.name_of_thm thm, thm);  | 
| 
4022
 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 
wenzelm 
parents: 
4013 
diff
changeset
 | 
147  | 
|
| 
 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 
wenzelm 
parents: 
4013 
diff
changeset
 | 
148  | 
fun thms_of thy =  | 
| 
 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 
wenzelm 
parents: 
4013 
diff
changeset
 | 
149  | 
  let val ref {thms_tab, ...} = get_theorems thy in
 | 
| 
 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 
wenzelm 
parents: 
4013 
diff
changeset
 | 
150  | 
map attach_name (flat (map snd (Symtab.dest thms_tab)))  | 
| 
 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 
wenzelm 
parents: 
4013 
diff
changeset
 | 
151  | 
end;  | 
| 
 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 
wenzelm 
parents: 
4013 
diff
changeset
 | 
152  | 
|
| 3987 | 153  | 
|
154  | 
||
| 
4022
 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 
wenzelm 
parents: 
4013 
diff
changeset
 | 
155  | 
(** theorems indexed by constants **)  | 
| 
 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 
wenzelm 
parents: 
4013 
diff
changeset
 | 
156  | 
|
| 
 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 
wenzelm 
parents: 
4013 
diff
changeset
 | 
157  | 
(* make index *)  | 
| 
 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 
wenzelm 
parents: 
4013 
diff
changeset
 | 
158  | 
|
| 4037 | 159  | 
val ignore = ["Trueprop", "all", "==>", "=="];  | 
| 
4022
 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 
wenzelm 
parents: 
4013 
diff
changeset
 | 
160  | 
|
| 6091 | 161  | 
fun add_const_idx ((next, table), thm) =  | 
| 
4022
 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 
wenzelm 
parents: 
4013 
diff
changeset
 | 
162  | 
let  | 
| 
 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 
wenzelm 
parents: 
4013 
diff
changeset
 | 
163  | 
    val {hyps, prop, ...} = Thm.rep_thm thm;
 | 
| 
 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 
wenzelm 
parents: 
4013 
diff
changeset
 | 
164  | 
val consts =  | 
| 4037 | 165  | 
foldr add_term_consts (hyps, add_term_consts (prop, [])) \\ ignore;  | 
| 
4022
 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 
wenzelm 
parents: 
4013 
diff
changeset
 | 
166  | 
|
| 
 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 
wenzelm 
parents: 
4013 
diff
changeset
 | 
167  | 
fun add (tab, c) =  | 
| 6091 | 168  | 
Symtab.update ((c, (next, thm) :: Symtab.lookup_multi (tab, c)), tab);  | 
| 
4022
 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 
wenzelm 
parents: 
4013 
diff
changeset
 | 
169  | 
in (next + 1, foldl add (table, consts)) end;  | 
| 
 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 
wenzelm 
parents: 
4013 
diff
changeset
 | 
170  | 
|
| 
 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 
wenzelm 
parents: 
4013 
diff
changeset
 | 
171  | 
fun make_const_idx thm_tab =  | 
| 5686 | 172  | 
Symtab.foldl (fn (x, (_, ths)) => foldl add_const_idx (x, ths)) ((0, Symtab.empty), thm_tab);  | 
| 
4022
 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 
wenzelm 
parents: 
4013 
diff
changeset
 | 
173  | 
|
| 
 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 
wenzelm 
parents: 
4013 
diff
changeset
 | 
174  | 
|
| 
 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 
wenzelm 
parents: 
4013 
diff
changeset
 | 
175  | 
(* lookup index *)  | 
| 3987 | 176  | 
|
| 
4022
 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 
wenzelm 
parents: 
4013 
diff
changeset
 | 
177  | 
(*search locally*)  | 
| 
 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 
wenzelm 
parents: 
4013 
diff
changeset
 | 
178  | 
fun containing [] thy = thms_of thy  | 
| 
 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 
wenzelm 
parents: 
4013 
diff
changeset
 | 
179  | 
| containing consts thy =  | 
| 
 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 
wenzelm 
parents: 
4013 
diff
changeset
 | 
180  | 
let  | 
| 4037 | 181  | 
fun int ([], _) = []  | 
182  | 
| int (_, []) = []  | 
|
183  | 
| int (xxs as ((x as (i:int, _)) :: xs), yys as ((y as (j, _)) :: ys)) =  | 
|
184  | 
if i = j then x :: int (xs, ys)  | 
|
185  | 
else if i > j then int (xs, yys)  | 
|
186  | 
else int (xxs, ys);  | 
|
| 
4022
 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 
wenzelm 
parents: 
4013 
diff
changeset
 | 
187  | 
|
| 4037 | 188  | 
fun ints [xs] = xs  | 
189  | 
| ints xss = if exists null xss then [] else foldl int (hd xss, tl xss);  | 
|
| 
4022
 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 
wenzelm 
parents: 
4013 
diff
changeset
 | 
190  | 
|
| 
 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 
wenzelm 
parents: 
4013 
diff
changeset
 | 
191  | 
        val ref {const_idx = (_, ctab), ...} = get_theorems thy;
 | 
| 
 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 
wenzelm 
parents: 
4013 
diff
changeset
 | 
192  | 
val ithmss = map (fn c => Symtab.lookup_multi (ctab, c)) consts;  | 
| 
 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 
wenzelm 
parents: 
4013 
diff
changeset
 | 
193  | 
in  | 
| 4037 | 194  | 
map (attach_name o snd) (ints ithmss)  | 
| 
4022
 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 
wenzelm 
parents: 
4013 
diff
changeset
 | 
195  | 
end;  | 
| 
 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 
wenzelm 
parents: 
4013 
diff
changeset
 | 
196  | 
|
| 
 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 
wenzelm 
parents: 
4013 
diff
changeset
 | 
197  | 
(*search globally*)  | 
| 
 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 
wenzelm 
parents: 
4013 
diff
changeset
 | 
198  | 
fun thms_containing thy consts =  | 
| 4037 | 199  | 
flat (map (containing (consts \\ ignore)) (thy :: Theory.ancestors_of thy));  | 
| 
4022
 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 
wenzelm 
parents: 
4013 
diff
changeset
 | 
200  | 
|
| 
 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 
wenzelm 
parents: 
4013 
diff
changeset
 | 
201  | 
|
| 
 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 
wenzelm 
parents: 
4013 
diff
changeset
 | 
202  | 
|
| 
 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 
wenzelm 
parents: 
4013 
diff
changeset
 | 
203  | 
(** store theorems **) (*DESTRUCTIVE*)  | 
| 3987 | 204  | 
|
| 4853 | 205  | 
(* naming *)  | 
206  | 
||
| 5933 | 207  | 
val defaultN = "it";  | 
208  | 
val default_name = fn "" => defaultN | name => name;  | 
|
| 5907 | 209  | 
|
| 4853 | 210  | 
fun gen_names len name =  | 
211  | 
map (fn i => name ^ "_" ^ string_of_int i) (1 upto len);  | 
|
212  | 
||
| 5907 | 213  | 
fun name_single name x = [(default_name name, x)];  | 
214  | 
fun name_multi name xs = gen_names (length xs) (default_name name) ~~ xs;  | 
|
| 4853 | 215  | 
|
216  | 
||
| 6091 | 217  | 
(* enter_thmx *)  | 
| 4853 | 218  | 
|
| 5933 | 219  | 
fun cond_warning name msg = if Sign.base_name name = defaultN then () else warning msg;  | 
220  | 
||
| 3987 | 221  | 
fun warn_overwrite name =  | 
| 5933 | 222  | 
  cond_warning name ("Replaced old copy of theorems " ^ quote name);
 | 
| 3987 | 223  | 
|
224  | 
fun warn_same name =  | 
|
| 5933 | 225  | 
  cond_warning name ("Theorem database already contains a copy of " ^ quote name);
 | 
| 3987 | 226  | 
|
| 6091 | 227  | 
fun enter_thmx sg app_name (bname, thmx) =  | 
| 3987 | 228  | 
let  | 
| 6350 | 229  | 
val name = Sign.full_name sg (default_name bname);  | 
| 6091 | 230  | 
val named_thms = map Thm.name_thm (app_name name thmx);  | 
| 
4022
 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 
wenzelm 
parents: 
4013 
diff
changeset
 | 
231  | 
|
| 
 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 
wenzelm 
parents: 
4013 
diff
changeset
 | 
232  | 
    val r as ref {space, thms_tab, const_idx} = get_theorems_sg sg;
 | 
| 
 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 
wenzelm 
parents: 
4013 
diff
changeset
 | 
233  | 
|
| 
 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 
wenzelm 
parents: 
4013 
diff
changeset
 | 
234  | 
val overwrite =  | 
| 
 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 
wenzelm 
parents: 
4013 
diff
changeset
 | 
235  | 
(case Symtab.lookup (thms_tab, name) of  | 
| 
 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 
wenzelm 
parents: 
4013 
diff
changeset
 | 
236  | 
None => false  | 
| 6091 | 237  | 
| Some thms' =>  | 
238  | 
if length thms' = length named_thms andalso forall2 Thm.eq_thm (thms', named_thms) then  | 
|
| 
4022
 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 
wenzelm 
parents: 
4013 
diff
changeset
 | 
239  | 
(warn_same name; false)  | 
| 
 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 
wenzelm 
parents: 
4013 
diff
changeset
 | 
240  | 
else (warn_overwrite name; true));  | 
| 
 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 
wenzelm 
parents: 
4013 
diff
changeset
 | 
241  | 
|
| 4487 | 242  | 
val space' = NameSpace.extend (space, [name]);  | 
| 6091 | 243  | 
val thms_tab' = Symtab.update ((name, named_thms), thms_tab);  | 
| 
4022
 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 
wenzelm 
parents: 
4013 
diff
changeset
 | 
244  | 
val const_idx' =  | 
| 
 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 
wenzelm 
parents: 
4013 
diff
changeset
 | 
245  | 
if overwrite then make_const_idx thms_tab'  | 
| 6091 | 246  | 
else foldl add_const_idx (const_idx, named_thms);  | 
| 3987 | 247  | 
in  | 
| 
4022
 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 
wenzelm 
parents: 
4013 
diff
changeset
 | 
248  | 
    r := {space = space', thms_tab = thms_tab', const_idx = const_idx'};
 | 
| 6091 | 249  | 
named_thms  | 
| 3987 | 250  | 
end;  | 
251  | 
||
| 4853 | 252  | 
|
| 6091 | 253  | 
(* add_thms(s) *)  | 
| 4853 | 254  | 
|
| 6091 | 255  | 
fun add_thmx app_name app_att ((bname, thmx), atts) thy =  | 
| 5280 | 256  | 
let  | 
| 6091 | 257  | 
val (thy', thmx') = app_att ((thy, thmx), atts);  | 
258  | 
val thms'' = enter_thmx (Theory.sign_of thy') app_name (bname, thmx');  | 
|
259  | 
in (thy', thms'') end;  | 
|
| 4853 | 260  | 
|
| 6091 | 261  | 
fun add_thmxs name app = Library.apply o map (fst oo add_thmx name app);  | 
| 5907 | 262  | 
|
| 6091 | 263  | 
val add_thms = add_thmxs name_single Thm.apply_attributes;  | 
264  | 
val add_thmss = add_thmxs name_multi Thm.applys_attributes;  | 
|
| 5907 | 265  | 
|
266  | 
||
| 6091 | 267  | 
(* have_thmss *)  | 
| 5907 | 268  | 
|
| 6367 | 269  | 
fun have_thmss opt_bname more_atts ths_atts thy =  | 
| 5907 | 270  | 
let  | 
| 6091 | 271  | 
fun app (x, (ths, atts)) = Thm.applys_attributes ((x, ths), atts);  | 
272  | 
val (thy', thmss') =  | 
|
| 5907 | 273  | 
foldl_map app (thy, map (fn (ths, atts) => (ths, atts @ more_atts)) ths_atts);  | 
| 6367 | 274  | 
val thms' = flat thmss';  | 
275  | 
val thms'' =  | 
|
276  | 
(case opt_bname of  | 
|
277  | 
None => thms'  | 
|
278  | 
| Some bname => enter_thmx (Theory.sign_of thy') name_multi (bname, thms'));  | 
|
| 6091 | 279  | 
in (thy, thms'') end;  | 
| 5280 | 280  | 
|
281  | 
||
| 6091 | 282  | 
(* store_thm *)  | 
| 5280 | 283  | 
|
| 6091 | 284  | 
fun store_thm th_atts thy =  | 
285  | 
let val (thy', [th']) = add_thmx name_single Thm.apply_attributes th_atts thy  | 
|
| 5280 | 286  | 
in (thy', th') end;  | 
| 3987 | 287  | 
|
288  | 
||
| 4853 | 289  | 
(* smart_store_thm *)  | 
| 3987 | 290  | 
|
| 4012 | 291  | 
fun smart_store_thm (name, thm) =  | 
| 6091 | 292  | 
hd (enter_thmx (Thm.sign_of_thm thm) name_single (name, thm));  | 
| 3987 | 293  | 
|
294  | 
||
| 
4022
 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 
wenzelm 
parents: 
4013 
diff
changeset
 | 
295  | 
(* store axioms as theorems *)  | 
| 
 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 
wenzelm 
parents: 
4013 
diff
changeset
 | 
296  | 
|
| 4853 | 297  | 
local  | 
298  | 
fun add_ax app_name add ((name, axs), atts) thy =  | 
|
299  | 
let  | 
|
300  | 
val named_axs = app_name name axs;  | 
|
301  | 
val thy' = add named_axs thy;  | 
|
| 6091 | 302  | 
val thms = map (Thm.get_axiom thy' o fst) named_axs;  | 
303  | 
in add_thmss [((name, thms), atts)] thy' end;  | 
|
| 
4022
 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 
wenzelm 
parents: 
4013 
diff
changeset
 | 
304  | 
|
| 5907 | 305  | 
fun add_axs app_name add = Library.apply o map (add_ax app_name add);  | 
| 4853 | 306  | 
in  | 
307  | 
val add_axioms = add_axs name_single Theory.add_axioms;  | 
|
308  | 
val add_axioms_i = add_axs name_single Theory.add_axioms_i;  | 
|
309  | 
val add_axiomss = add_axs name_multi Theory.add_axioms;  | 
|
310  | 
val add_axiomss_i = add_axs name_multi Theory.add_axioms_i;  | 
|
311  | 
val add_defs = add_axs name_single Theory.add_defs;  | 
|
312  | 
val add_defs_i = add_axs name_single Theory.add_defs_i;  | 
|
313  | 
val add_defss = add_axs name_multi Theory.add_defs;  | 
|
314  | 
val add_defss_i = add_axs name_multi Theory.add_defs_i;  | 
|
315  | 
end;  | 
|
| 
4022
 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 
wenzelm 
parents: 
4013 
diff
changeset
 | 
316  | 
|
| 
 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 
wenzelm 
parents: 
4013 
diff
changeset
 | 
317  | 
|
| 3987 | 318  | 
|
| 
4963
 
38aa2d56e28c
added get_name, put_name, global_path, local_path, begin_theory,
 
wenzelm 
parents: 
4933 
diff
changeset
 | 
319  | 
(*** derived theory operations ***)  | 
| 
 
38aa2d56e28c
added get_name, put_name, global_path, local_path, begin_theory,
 
wenzelm 
parents: 
4933 
diff
changeset
 | 
320  | 
|
| 
 
38aa2d56e28c
added get_name, put_name, global_path, local_path, begin_theory,
 
wenzelm 
parents: 
4933 
diff
changeset
 | 
321  | 
(** theory management **)  | 
| 
 
38aa2d56e28c
added get_name, put_name, global_path, local_path, begin_theory,
 
wenzelm 
parents: 
4933 
diff
changeset
 | 
322  | 
|
| 5005 | 323  | 
(* data kind 'Pure/theory_management' *)  | 
| 
4963
 
38aa2d56e28c
added get_name, put_name, global_path, local_path, begin_theory,
 
wenzelm 
parents: 
4933 
diff
changeset
 | 
324  | 
|
| 5005 | 325  | 
structure TheoryManagementDataArgs =  | 
326  | 
struct  | 
|
327  | 
val name = "Pure/theory_management";  | 
|
328  | 
  type T = {name: string, generation: int};
 | 
|
| 5000 | 329  | 
|
| 5005 | 330  | 
  val empty = {name = "", generation = 0};
 | 
| 6547 | 331  | 
val copy = I;  | 
| 5005 | 332  | 
val prep_ext = I;  | 
| 5000 | 333  | 
fun merge _ = empty;  | 
| 5005 | 334  | 
fun print _ _ = ();  | 
| 
4963
 
38aa2d56e28c
added get_name, put_name, global_path, local_path, begin_theory,
 
wenzelm 
parents: 
4933 
diff
changeset
 | 
335  | 
end;  | 
| 
 
38aa2d56e28c
added get_name, put_name, global_path, local_path, begin_theory,
 
wenzelm 
parents: 
4933 
diff
changeset
 | 
336  | 
|
| 5005 | 337  | 
structure TheoryManagementData = TheoryDataFun(TheoryManagementDataArgs);  | 
338  | 
val get_info = TheoryManagementData.get;  | 
|
339  | 
val put_info = TheoryManagementData.put;  | 
|
340  | 
||
| 
4963
 
38aa2d56e28c
added get_name, put_name, global_path, local_path, begin_theory,
 
wenzelm 
parents: 
4933 
diff
changeset
 | 
341  | 
|
| 
 
38aa2d56e28c
added get_name, put_name, global_path, local_path, begin_theory,
 
wenzelm 
parents: 
4933 
diff
changeset
 | 
342  | 
(* get / put name *)  | 
| 
 
38aa2d56e28c
added get_name, put_name, global_path, local_path, begin_theory,
 
wenzelm 
parents: 
4933 
diff
changeset
 | 
343  | 
|
| 5000 | 344  | 
val get_name = #name o get_info;  | 
345  | 
fun put_name name = put_info {name = name, generation = 0};
 | 
|
| 
4963
 
38aa2d56e28c
added get_name, put_name, global_path, local_path, begin_theory,
 
wenzelm 
parents: 
4933 
diff
changeset
 | 
346  | 
|
| 
 
38aa2d56e28c
added get_name, put_name, global_path, local_path, begin_theory,
 
wenzelm 
parents: 
4933 
diff
changeset
 | 
347  | 
|
| 
 
38aa2d56e28c
added get_name, put_name, global_path, local_path, begin_theory,
 
wenzelm 
parents: 
4933 
diff
changeset
 | 
348  | 
(* control prefixing of theory name *)  | 
| 
 
38aa2d56e28c
added get_name, put_name, global_path, local_path, begin_theory,
 
wenzelm 
parents: 
4933 
diff
changeset
 | 
349  | 
|
| 5210 | 350  | 
val global_path = Theory.root_path;  | 
| 
4963
 
38aa2d56e28c
added get_name, put_name, global_path, local_path, begin_theory,
 
wenzelm 
parents: 
4933 
diff
changeset
 | 
351  | 
|
| 
 
38aa2d56e28c
added get_name, put_name, global_path, local_path, begin_theory,
 
wenzelm 
parents: 
4933 
diff
changeset
 | 
352  | 
fun local_path thy =  | 
| 5210 | 353  | 
thy |> Theory.root_path |> Theory.add_path (get_name thy);  | 
| 
4963
 
38aa2d56e28c
added get_name, put_name, global_path, local_path, begin_theory,
 
wenzelm 
parents: 
4933 
diff
changeset
 | 
354  | 
|
| 
 
38aa2d56e28c
added get_name, put_name, global_path, local_path, begin_theory,
 
wenzelm 
parents: 
4933 
diff
changeset
 | 
355  | 
|
| 
 
38aa2d56e28c
added get_name, put_name, global_path, local_path, begin_theory,
 
wenzelm 
parents: 
4933 
diff
changeset
 | 
356  | 
(* begin / end theory *)  | 
| 
 
38aa2d56e28c
added get_name, put_name, global_path, local_path, begin_theory,
 
wenzelm 
parents: 
4933 
diff
changeset
 | 
357  | 
|
| 
 
38aa2d56e28c
added get_name, put_name, global_path, local_path, begin_theory,
 
wenzelm 
parents: 
4933 
diff
changeset
 | 
358  | 
fun begin_theory name thys =  | 
| 
 
38aa2d56e28c
added get_name, put_name, global_path, local_path, begin_theory,
 
wenzelm 
parents: 
4933 
diff
changeset
 | 
359  | 
Theory.prep_ext_merge thys  | 
| 
 
38aa2d56e28c
added get_name, put_name, global_path, local_path, begin_theory,
 
wenzelm 
parents: 
4933 
diff
changeset
 | 
360  | 
|> put_name name  | 
| 
 
38aa2d56e28c
added get_name, put_name, global_path, local_path, begin_theory,
 
wenzelm 
parents: 
4933 
diff
changeset
 | 
361  | 
|> local_path;  | 
| 
 
38aa2d56e28c
added get_name, put_name, global_path, local_path, begin_theory,
 
wenzelm 
parents: 
4933 
diff
changeset
 | 
362  | 
|
| 
 
38aa2d56e28c
added get_name, put_name, global_path, local_path, begin_theory,
 
wenzelm 
parents: 
4933 
diff
changeset
 | 
363  | 
fun end_theory thy = Theory.add_name (get_name thy) thy;  | 
| 
 
38aa2d56e28c
added get_name, put_name, global_path, local_path, begin_theory,
 
wenzelm 
parents: 
4933 
diff
changeset
 | 
364  | 
|
| 
 
38aa2d56e28c
added get_name, put_name, global_path, local_path, begin_theory,
 
wenzelm 
parents: 
4933 
diff
changeset
 | 
365  | 
|
| 5000 | 366  | 
(* atomic transactions *)  | 
367  | 
||
| 5091 | 368  | 
exception ROLLBACK of theory * exn option;  | 
369  | 
||
| 5000 | 370  | 
fun transaction f thy =  | 
371  | 
let  | 
|
372  | 
    val {name, generation} = get_info thy;
 | 
|
373  | 
val copy_thy =  | 
|
374  | 
thy  | 
|
375  | 
|> Theory.prep_ext  | 
|
| 5005 | 376  | 
      |> Theory.add_name ("#" ^ name ^ ":" ^ string_of_int generation)  (* FIXME !!?? *)
 | 
| 5000 | 377  | 
      |> put_info {name = name, generation = generation + 1};
 | 
| 5026 | 378  | 
val (thy', opt_exn) = (transform_error f thy, None) handle exn => (thy, Some exn);  | 
| 5000 | 379  | 
in  | 
| 5091 | 380  | 
if Sign.is_stale (Theory.sign_of thy') then raise ROLLBACK (copy_thy, opt_exn)  | 
381  | 
else (case opt_exn of Some exn => raise exn | _ => thy')  | 
|
| 5000 | 382  | 
end;  | 
383  | 
||
384  | 
||
| 
4963
 
38aa2d56e28c
added get_name, put_name, global_path, local_path, begin_theory,
 
wenzelm 
parents: 
4933 
diff
changeset
 | 
385  | 
|
| 
 
38aa2d56e28c
added get_name, put_name, global_path, local_path, begin_theory,
 
wenzelm 
parents: 
4933 
diff
changeset
 | 
386  | 
(** add logical types **)  | 
| 
4922
 
03b81b6e1baa
added thms_closure: theory -> xstring -> tthm list option;
 
wenzelm 
parents: 
4853 
diff
changeset
 | 
387  | 
|
| 
 
03b81b6e1baa
added thms_closure: theory -> xstring -> tthm list option;
 
wenzelm 
parents: 
4853 
diff
changeset
 | 
388  | 
fun add_typedecls decls thy =  | 
| 
 
03b81b6e1baa
added thms_closure: theory -> xstring -> tthm list option;
 
wenzelm 
parents: 
4853 
diff
changeset
 | 
389  | 
let  | 
| 
 
03b81b6e1baa
added thms_closure: theory -> xstring -> tthm list option;
 
wenzelm 
parents: 
4853 
diff
changeset
 | 
390  | 
val full = Sign.full_name (Theory.sign_of thy);  | 
| 
 
03b81b6e1baa
added thms_closure: theory -> xstring -> tthm list option;
 
wenzelm 
parents: 
4853 
diff
changeset
 | 
391  | 
|
| 
 
03b81b6e1baa
added thms_closure: theory -> xstring -> tthm list option;
 
wenzelm 
parents: 
4853 
diff
changeset
 | 
392  | 
fun type_of (raw_name, vs, mx) =  | 
| 
 
03b81b6e1baa
added thms_closure: theory -> xstring -> tthm list option;
 
wenzelm 
parents: 
4853 
diff
changeset
 | 
393  | 
if null (duplicates vs) then (raw_name, length vs, mx)  | 
| 
 
03b81b6e1baa
added thms_closure: theory -> xstring -> tthm list option;
 
wenzelm 
parents: 
4853 
diff
changeset
 | 
394  | 
      else error ("Duplicate parameters in type declaration: " ^ quote raw_name);
 | 
| 
 
03b81b6e1baa
added thms_closure: theory -> xstring -> tthm list option;
 
wenzelm 
parents: 
4853 
diff
changeset
 | 
395  | 
|
| 
 
03b81b6e1baa
added thms_closure: theory -> xstring -> tthm list option;
 
wenzelm 
parents: 
4853 
diff
changeset
 | 
396  | 
fun arity_of (raw_name, len, mx) =  | 
| 
 
03b81b6e1baa
added thms_closure: theory -> xstring -> tthm list option;
 
wenzelm 
parents: 
4853 
diff
changeset
 | 
397  | 
(full (Syntax.type_name raw_name mx), replicate len logicS, logicS);  | 
| 
 
03b81b6e1baa
added thms_closure: theory -> xstring -> tthm list option;
 
wenzelm 
parents: 
4853 
diff
changeset
 | 
398  | 
|
| 
 
03b81b6e1baa
added thms_closure: theory -> xstring -> tthm list option;
 
wenzelm 
parents: 
4853 
diff
changeset
 | 
399  | 
val types = map type_of decls;  | 
| 
 
03b81b6e1baa
added thms_closure: theory -> xstring -> tthm list option;
 
wenzelm 
parents: 
4853 
diff
changeset
 | 
400  | 
val arities = map arity_of types;  | 
| 
 
03b81b6e1baa
added thms_closure: theory -> xstring -> tthm list option;
 
wenzelm 
parents: 
4853 
diff
changeset
 | 
401  | 
in  | 
| 
 
03b81b6e1baa
added thms_closure: theory -> xstring -> tthm list option;
 
wenzelm 
parents: 
4853 
diff
changeset
 | 
402  | 
thy  | 
| 
 
03b81b6e1baa
added thms_closure: theory -> xstring -> tthm list option;
 
wenzelm 
parents: 
4853 
diff
changeset
 | 
403  | 
|> Theory.add_types types  | 
| 
 
03b81b6e1baa
added thms_closure: theory -> xstring -> tthm list option;
 
wenzelm 
parents: 
4853 
diff
changeset
 | 
404  | 
|> Theory.add_arities_i arities  | 
| 
 
03b81b6e1baa
added thms_closure: theory -> xstring -> tthm list option;
 
wenzelm 
parents: 
4853 
diff
changeset
 | 
405  | 
end;  | 
| 
 
03b81b6e1baa
added thms_closure: theory -> xstring -> tthm list option;
 
wenzelm 
parents: 
4853 
diff
changeset
 | 
406  | 
|
| 
 
03b81b6e1baa
added thms_closure: theory -> xstring -> tthm list option;
 
wenzelm 
parents: 
4853 
diff
changeset
 | 
407  | 
|
| 
 
03b81b6e1baa
added thms_closure: theory -> xstring -> tthm list option;
 
wenzelm 
parents: 
4853 
diff
changeset
 | 
408  | 
|
| 5091 | 409  | 
(*** the ProtoPure theory ***)  | 
| 3987 | 410  | 
|
| 
6560
 
1436349f8b28
renamed 'dummy' to 'dummy_pattern' (less dangerous);
 
wenzelm 
parents: 
6547 
diff
changeset
 | 
411  | 
val dummy_patternN = "dummy_pattern";  | 
| 
 
1436349f8b28
renamed 'dummy' to 'dummy_pattern' (less dangerous);
 
wenzelm 
parents: 
6547 
diff
changeset
 | 
412  | 
|
| 3987 | 413  | 
val proto_pure =  | 
414  | 
Theory.pre_pure  | 
|
| 5907 | 415  | 
|> Library.apply [TheoremsData.init, TheoryManagementData.init]  | 
| 
4963
 
38aa2d56e28c
added get_name, put_name, global_path, local_path, begin_theory,
 
wenzelm 
parents: 
4933 
diff
changeset
 | 
416  | 
|> put_name "ProtoPure"  | 
| 
 
38aa2d56e28c
added get_name, put_name, global_path, local_path, begin_theory,
 
wenzelm 
parents: 
4933 
diff
changeset
 | 
417  | 
|> global_path  | 
| 3987 | 418  | 
|> Theory.add_types  | 
| 
4922
 
03b81b6e1baa
added thms_closure: theory -> xstring -> tthm list option;
 
wenzelm 
parents: 
4853 
diff
changeset
 | 
419  | 
   [("fun", 2, NoSyn),
 | 
| 
 
03b81b6e1baa
added thms_closure: theory -> xstring -> tthm list option;
 
wenzelm 
parents: 
4853 
diff
changeset
 | 
420  | 
    ("prop", 0, NoSyn),
 | 
| 
 
03b81b6e1baa
added thms_closure: theory -> xstring -> tthm list option;
 
wenzelm 
parents: 
4853 
diff
changeset
 | 
421  | 
    ("itself", 1, NoSyn),
 | 
| 
 
03b81b6e1baa
added thms_closure: theory -> xstring -> tthm list option;
 
wenzelm 
parents: 
4853 
diff
changeset
 | 
422  | 
    ("dummy", 0, NoSyn)]
 | 
| 3987 | 423  | 
|> Theory.add_classes_i [(logicC, [])]  | 
424  | 
|> Theory.add_defsort_i logicS  | 
|
425  | 
|> Theory.add_arities_i  | 
|
426  | 
   [("fun", [logicS, logicS], logicS),
 | 
|
427  | 
    ("prop", [], logicS),
 | 
|
428  | 
    ("itself", [logicS], logicS)]
 | 
|
| 
4922
 
03b81b6e1baa
added thms_closure: theory -> xstring -> tthm list option;
 
wenzelm 
parents: 
4853 
diff
changeset
 | 
429  | 
|> Theory.add_nonterminals Syntax.pure_nonterms  | 
| 3987 | 430  | 
|> Theory.add_syntax Syntax.pure_syntax  | 
| 
6027
 
9dd06eeda95c
added new print_mode "xsymbols" for extended symbol support
 
oheimb 
parents: 
5933 
diff
changeset
 | 
431  | 
  |> Theory.add_modesyntax ("symbols" , true) Syntax.pure_sym_syntax
 | 
| 
 
9dd06eeda95c
added new print_mode "xsymbols" for extended symbol support
 
oheimb 
parents: 
5933 
diff
changeset
 | 
432  | 
  |> Theory.add_modesyntax ("xsymbols", true) Syntax.pure_xsym_syntax
 | 
| 3987 | 433  | 
|> Theory.add_trfuns Syntax.pure_trfuns  | 
434  | 
|> Theory.add_trfunsT Syntax.pure_trfunsT  | 
|
435  | 
|> Theory.add_syntax  | 
|
436  | 
   [("==>", "[prop, prop] => prop", Delimfix "op ==>")]
 | 
|
437  | 
|> Theory.add_consts  | 
|
438  | 
   [("==", "['a::{}, 'a] => prop", InfixrName ("==", 2)),
 | 
|
439  | 
    ("=?=", "['a::{}, 'a] => prop", InfixrName ("=?=", 2)),
 | 
|
440  | 
    ("==>", "[prop, prop] => prop", Mixfix ("(_/ ==> _)", [2, 1], 1)),
 | 
|
441  | 
    ("all", "('a => prop) => prop", Binder ("!!", 0, 0)),
 | 
|
| 4788 | 442  | 
    ("Goal", "prop => prop", Mixfix ("GOAL _", [999], 1000)),
 | 
| 6547 | 443  | 
    ("TYPE", "'a itself", NoSyn),
 | 
| 
6560
 
1436349f8b28
renamed 'dummy' to 'dummy_pattern' (less dangerous);
 
wenzelm 
parents: 
6547 
diff
changeset
 | 
444  | 
(dummy_patternN, "'a", Delimfix "'_")]  | 
| 
5041
 
a1d0a6d555cd
Goals may now contain assumptions, which are not returned.
 
nipkow 
parents: 
5026 
diff
changeset
 | 
445  | 
  |> Theory.add_modesyntax ("", false)
 | 
| 
 
a1d0a6d555cd
Goals may now contain assumptions, which are not returned.
 
nipkow 
parents: 
5026 
diff
changeset
 | 
446  | 
    [("Goal", "prop => prop", Mixfix ("_", [0], 0))]
 | 
| 
4963
 
38aa2d56e28c
added get_name, put_name, global_path, local_path, begin_theory,
 
wenzelm 
parents: 
4933 
diff
changeset
 | 
447  | 
|> local_path  | 
| 6091 | 448  | 
|> (add_defs o map Thm.no_attributes)  | 
| 4788 | 449  | 
   [("flexpair_def", "(t =?= u) == (t == u::'a::{})"),
 | 
450  | 
    ("Goal_def", "GOAL (PROP A) == PROP A")]
 | 
|
| 
4963
 
38aa2d56e28c
added get_name, put_name, global_path, local_path, begin_theory,
 
wenzelm 
parents: 
4933 
diff
changeset
 | 
451  | 
|> end_theory;  | 
| 3987 | 452  | 
|
| 5091 | 453  | 
structure ProtoPure =  | 
454  | 
struct  | 
|
455  | 
val thy = proto_pure;  | 
|
456  | 
val flexpair_def = get_axiom thy "flexpair_def";  | 
|
457  | 
val Goal_def = get_axiom thy "Goal_def";  | 
|
458  | 
end;  | 
|
| 3987 | 459  | 
|
460  | 
||
461  | 
end;  | 
|
462  | 
||
463  | 
||
| 
4022
 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 
wenzelm 
parents: 
4013 
diff
changeset
 | 
464  | 
structure BasicPureThy: BASIC_PURE_THY = PureThy;  | 
| 
 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 
wenzelm 
parents: 
4013 
diff
changeset
 | 
465  | 
open BasicPureThy;  |