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