| author | wenzelm | 
| Wed, 08 Jun 2011 20:58:51 +0200 | |
| changeset 43284 | 04d473e883df | 
| parent 42471 | 593289343c7d | 
| child 44304 | 7ee000ce5390 | 
| permissions | -rw-r--r-- | 
| 39557 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 wenzelm parents: 
39507diff
changeset | 1 | (* Title: Pure/global_theory.ML | 
| 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 wenzelm parents: 
39507diff
changeset | 2 | Author: Makarius | 
| 3987 | 3 | |
| 39557 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 wenzelm parents: 
39507diff
changeset | 4 | Global theory content: stored facts. | 
| 3987 | 5 | *) | 
| 6 | ||
| 39557 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 wenzelm parents: 
39507diff
changeset | 7 | signature GLOBAL_THEORY = | 
| 3987 | 8 | sig | 
| 27198 | 9 | val facts_of: theory -> Facts.T | 
| 26666 | 10 | val intern_fact: theory -> xstring -> string | 
| 26693 | 11 | val defined_fact: theory -> string -> bool | 
| 27198 | 12 | val hide_fact: bool -> string -> theory -> theory | 
| 32105 | 13 | val join_proofs: theory -> unit | 
| 26683 | 14 | val get_fact: Context.generic -> theory -> Facts.ref -> thm list | 
| 26344 
04dacc6809b6
simplified get_thm(s): back to plain name argument;
 wenzelm parents: 
26336diff
changeset | 15 | val get_thms: theory -> xstring -> thm list | 
| 
04dacc6809b6
simplified get_thm(s): back to plain name argument;
 wenzelm parents: 
26336diff
changeset | 16 | val get_thm: theory -> xstring -> thm | 
| 16336 | 17 | val all_thms_of: theory -> (string * thm) list | 
| 21580 | 18 |   val map_facts: ('a -> 'b) -> ('c * ('a list * 'd) list) list -> ('c * ('b list * 'd) list) list
 | 
| 21567 | 19 |   val burrow_fact: ('a list -> 'b list) -> ('a list * 'c) list -> ('b list * 'c) list
 | 
| 21580 | 20 |   val burrow_facts: ('a list -> 'b list) ->
 | 
| 21 |     ('c * ('a list * 'd) list) list -> ('c * ('b list * 'd) list) list
 | |
| 22 | val name_multi: string -> 'a list -> (string * 'a) list | |
| 33700 | 23 | val name_thm: bool -> bool -> string -> thm -> thm | 
| 24 | val name_thms: bool -> bool -> string -> thm list -> thm list | |
| 25 | val name_thmss: bool -> string -> (thm list * 'a) list -> (thm list * 'a) list | |
| 29579 | 26 | val store_thms: binding * thm list -> theory -> thm list * theory | 
| 27 | val store_thm: binding * thm -> theory -> thm * theory | |
| 28 | val store_thm_open: binding * thm -> theory -> thm * theory | |
| 29 | val add_thms: ((binding * thm) * attribute list) list -> theory -> thm list * theory | |
| 30 | val add_thm: (binding * thm) * attribute list -> theory -> thm * theory | |
| 31 | val add_thmss: ((binding * thm list) * attribute list) list -> theory -> thm list list * theory | |
| 32 | val add_thms_dynamic: binding * (Context.generic -> thm list) -> theory -> theory | |
| 30853 | 33 | val note_thmss: string -> (Thm.binding * (thm list * attribute list) list) list | 
| 34 | -> theory -> (string * thm list) list * theory | |
| 29579 | 35 | val add_defs: bool -> ((binding * term) * attribute list) list -> | 
| 18377 | 36 | theory -> thm list * theory | 
| 29579 | 37 | val add_defs_unchecked: bool -> ((binding * term) * attribute list) list -> | 
| 38 | theory -> thm list * theory | |
| 30337 
eb189f7e43a1
Theory.add_axioms/add_defs: replaced old bstring by binding;
 wenzelm parents: 
30242diff
changeset | 39 | val add_defs_cmd: bool -> ((binding * string) * attribute list) list -> | 
| 18377 | 40 | theory -> thm list * theory | 
| 30337 
eb189f7e43a1
Theory.add_axioms/add_defs: replaced old bstring by binding;
 wenzelm parents: 
30242diff
changeset | 41 | val add_defs_unchecked_cmd: bool -> ((binding * string) * attribute list) list -> | 
| 19629 | 42 | theory -> thm list * theory | 
| 3987 | 43 | end; | 
| 44 | ||
| 39557 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 wenzelm parents: 
39507diff
changeset | 45 | structure Global_Theory: GLOBAL_THEORY = | 
| 3987 | 46 | struct | 
| 47 | ||
| 27198 | 48 | (** theory data **) | 
| 26282 
305d5ca4fa9d
replaced obsolete FactIndex.T by Facts.T (cumulative version, assumes that facts are only added to unfinished theories);
 wenzelm parents: 
26050diff
changeset | 49 | |
| 39557 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 wenzelm parents: 
39507diff
changeset | 50 | structure Data = Theory_Data | 
| 24713 | 51 | ( | 
| 32105 | 52 | type T = Facts.T * thm list; | 
| 53 | val empty = (Facts.empty, []); | |
| 54 | fun extend (facts, _) = (facts, []); | |
| 33522 | 55 | fun merge ((facts1, _), (facts2, _)) = (Facts.merge (facts1, facts2), []); | 
| 24713 | 56 | ); | 
| 3987 | 57 | |
| 28841 
5556c7976837
added force_proofs (based on thms ever passed through enter_thms);
 wenzelm parents: 
28674diff
changeset | 58 | |
| 
5556c7976837
added force_proofs (based on thms ever passed through enter_thms);
 wenzelm parents: 
28674diff
changeset | 59 | (* facts *) | 
| 
5556c7976837
added force_proofs (based on thms ever passed through enter_thms);
 wenzelm parents: 
28674diff
changeset | 60 | |
| 39557 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 wenzelm parents: 
39507diff
changeset | 61 | val facts_of = #1 o Data.get; | 
| 26666 | 62 | |
| 63 | val intern_fact = Facts.intern o facts_of; | |
| 26693 | 64 | val defined_fact = Facts.defined o facts_of; | 
| 16023 
66561f6814bd
added string_of_thmref, selections, fact_index_of, valid_thms;
 wenzelm parents: 
15975diff
changeset | 65 | |
| 39557 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 wenzelm parents: 
39507diff
changeset | 66 | fun hide_fact fully name = Data.map (apfst (Facts.hide fully name)); | 
| 28841 
5556c7976837
added force_proofs (based on thms ever passed through enter_thms);
 wenzelm parents: 
28674diff
changeset | 67 | |
| 
5556c7976837
added force_proofs (based on thms ever passed through enter_thms);
 wenzelm parents: 
28674diff
changeset | 68 | |
| 
5556c7976837
added force_proofs (based on thms ever passed through enter_thms);
 wenzelm parents: 
28674diff
changeset | 69 | (* proofs *) | 
| 
5556c7976837
added force_proofs (based on thms ever passed through enter_thms);
 wenzelm parents: 
28674diff
changeset | 70 | |
| 39557 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 wenzelm parents: 
39507diff
changeset | 71 | fun register_proofs (thy, thms) = (Data.map (apsnd (append thms)) thy, thms); | 
| 28977 | 72 | |
| 39557 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 wenzelm parents: 
39507diff
changeset | 73 | fun join_proofs thy = Thm.join_proofs (rev (#2 (Data.get thy))); | 
| 27198 | 74 | |
| 6367 | 75 | |
| 3987 | 76 | |
| 4022 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 wenzelm parents: 
4013diff
changeset | 77 | (** retrieve theorems **) | 
| 3987 | 78 | |
| 27198 | 79 | fun get_fact context thy xthmref = | 
| 26683 | 80 | let | 
| 42378 | 81 | val ctxt = Context.proof_of context; | 
| 82 | ||
| 83 | val facts = facts_of thy; | |
| 26683 | 84 | val xname = Facts.name_of_ref xthmref; | 
| 85 | val pos = Facts.pos_of_ref xthmref; | |
| 27198 | 86 | |
| 42471 
593289343c7d
hardwired mapping "_" -> "Pure.asm_rl" avoids legacy binding;
 wenzelm parents: 
42381diff
changeset | 87 | val name = | 
| 
593289343c7d
hardwired mapping "_" -> "Pure.asm_rl" avoids legacy binding;
 wenzelm parents: 
42381diff
changeset | 88 | (case intern_fact thy xname of | 
| 
593289343c7d
hardwired mapping "_" -> "Pure.asm_rl" avoids legacy binding;
 wenzelm parents: 
42381diff
changeset | 89 | "_" => "Pure.asm_rl" | 
| 
593289343c7d
hardwired mapping "_" -> "Pure.asm_rl" avoids legacy binding;
 wenzelm parents: 
42381diff
changeset | 90 | | name => name); | 
| 42378 | 91 | val res = Facts.lookup context facts name; | 
| 27198 | 92 | val _ = Theory.check_thy thy; | 
| 26683 | 93 | in | 
| 27198 | 94 | (case res of | 
| 95 |       NONE => error ("Unknown fact " ^ quote name ^ Position.str_of pos)
 | |
| 27739 | 96 | | SOME (static, ths) => | 
| 42379 | 97 | (Context_Position.report ctxt pos (Name_Space.markup (Facts.space_of facts) name); | 
| 42378 | 98 | if static then () else Context_Position.report ctxt pos (Markup.dynamic_fact name); | 
| 27739 | 99 | Facts.select xthmref (map (Thm.transfer thy) ths))) | 
| 26683 | 100 | end; | 
| 26344 
04dacc6809b6
simplified get_thm(s): back to plain name argument;
 wenzelm parents: 
26336diff
changeset | 101 | |
| 26683 | 102 | fun get_thms thy = get_fact (Context.Theory thy) thy o Facts.named; | 
| 26344 
04dacc6809b6
simplified get_thm(s): back to plain name argument;
 wenzelm parents: 
26336diff
changeset | 103 | fun get_thm thy name = Facts.the_single name (get_thms thy name); | 
| 4783 | 104 | |
| 27198 | 105 | fun all_thms_of thy = | 
| 27865 
27a8ad9612a3
moved basic thm operations from structure PureThy to Thm (cf. more_thm.ML);
 wenzelm parents: 
27739diff
changeset | 106 | Facts.fold_static (fn (_, ths) => append (map (`(Thm.get_name_hint)) ths)) (facts_of thy) []; | 
| 16336 | 107 | |
| 4022 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 wenzelm parents: 
4013diff
changeset | 108 | |
| 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 wenzelm parents: 
4013diff
changeset | 109 | |
| 26488 
b497e3187ec7
eliminated destructive/critical theorem database;
 wenzelm parents: 
26471diff
changeset | 110 | (** store theorems **) | 
| 3987 | 111 | |
| 21580 | 112 | (* fact specifications *) | 
| 113 | ||
| 114 | fun map_facts f = map (apsnd (map (apfst (map f)))); | |
| 115 | fun burrow_fact f = split_list #>> burrow f #> op ~~; | |
| 116 | fun burrow_facts f = split_list ##> burrow (burrow_fact f) #> op ~~; | |
| 117 | ||
| 118 | ||
| 4853 | 119 | (* naming *) | 
| 120 | ||
| 18801 | 121 | fun name_multi name [x] = [(name, x)] | 
| 26457 | 122 | | name_multi "" xs = map (pair "") xs | 
| 123 | | name_multi name xs = map_index (fn (i, x) => (name ^ "_" ^ string_of_int (i + 1), x)) xs; | |
| 12235 
5fa04fc9b254
Further restructuring of theorem naming functions.
 berghofe parents: 
12138diff
changeset | 124 | |
| 33700 | 125 | fun name_thm pre official name thm = thm | 
| 41696 | 126 | |> (if not official orelse pre andalso Thm.derivation_name thm <> "" then I | 
| 127 | else Thm.name_derivation name) | |
| 128 | |> (if name = "" orelse pre andalso Thm.has_name_hint thm then I | |
| 129 | else Thm.put_name_hint name); | |
| 12872 
0855c3ab2047
Theorems are only "pre-named" if the do not already have names.
 berghofe parents: 
12711diff
changeset | 130 | |
| 33700 | 131 | fun name_thms pre official name xs = | 
| 132 | map (uncurry (name_thm pre official)) (name_multi name xs); | |
| 12235 
5fa04fc9b254
Further restructuring of theorem naming functions.
 berghofe parents: 
12138diff
changeset | 133 | |
| 33700 | 134 | fun name_thmss official name fact = | 
| 135 | burrow_fact (name_thms true official name) fact; | |
| 4853 | 136 | |
| 137 | ||
| 11998 | 138 | (* enter_thms *) | 
| 4853 | 139 | |
| 28861 | 140 | fun enter_thms pre_name post_name app_att (b, thms) thy = | 
| 28965 | 141 | if Binding.is_empty b | 
| 32105 | 142 | then swap (register_proofs (app_att (thy, thms))) | 
| 30211 | 143 | else | 
| 144 | let | |
| 145 | val naming = Sign.naming_of thy; | |
| 33095 
bbd52d2f8696
renamed NameSpace to Name_Space -- also to emphasize its subtle change in semantics;
 wenzelm parents: 
32786diff
changeset | 146 | val name = Name_Space.full_name naming b; | 
| 30211 | 147 | val (thy', thms') = | 
| 32105 | 148 | register_proofs (apsnd (post_name name) (app_att (thy, pre_name name thms))); | 
| 30211 | 149 | val thms'' = map (Thm.transfer thy') thms'; | 
| 42375 
774df7c59508
report Name_Space.declare/define, relatively to context;
 wenzelm parents: 
41696diff
changeset | 150 | val thy'' = thy' | 
| 
774df7c59508
report Name_Space.declare/define, relatively to context;
 wenzelm parents: 
41696diff
changeset | 151 | |> (Data.map o apfst) | 
| 
774df7c59508
report Name_Space.declare/define, relatively to context;
 wenzelm parents: 
41696diff
changeset | 152 | (Facts.add_global (Proof_Context.init_global thy') naming (b, thms'') #> snd); | 
| 30211 | 153 | in (thms'', thy'') end; | 
| 26488 
b497e3187ec7
eliminated destructive/critical theorem database;
 wenzelm parents: 
26471diff
changeset | 154 | |
| 
b497e3187ec7
eliminated destructive/critical theorem database;
 wenzelm parents: 
26471diff
changeset | 155 | |
| 
b497e3187ec7
eliminated destructive/critical theorem database;
 wenzelm parents: 
26471diff
changeset | 156 | (* store_thm(s) *) | 
| 
b497e3187ec7
eliminated destructive/critical theorem database;
 wenzelm parents: 
26471diff
changeset | 157 | |
| 33700 | 158 | fun store_thms (b, thms) = | 
| 159 | enter_thms (name_thms true true) (name_thms false true) I (b, thms); | |
| 28076 | 160 | |
| 29579 | 161 | fun store_thm (b, th) = store_thms (b, [th]) #>> the_single; | 
| 26488 
b497e3187ec7
eliminated destructive/critical theorem database;
 wenzelm parents: 
26471diff
changeset | 162 | |
| 29579 | 163 | fun store_thm_open (b, th) = | 
| 33700 | 164 | enter_thms (name_thms true false) (name_thms false false) I (b, [th]) #>> the_single; | 
| 3987 | 165 | |
| 16023 
66561f6814bd
added string_of_thmref, selections, fact_index_of, valid_thms;
 wenzelm parents: 
15975diff
changeset | 166 | |
| 6091 | 167 | (* add_thms(s) *) | 
| 4853 | 168 | |
| 29579 | 169 | fun add_thms_atts pre_name ((b, thms), atts) = | 
| 33700 | 170 | enter_thms pre_name (name_thms false true) | 
| 30190 | 171 | (Library.foldl_map (Thm.theory_attributes atts)) (b, thms); | 
| 4853 | 172 | |
| 18377 | 173 | fun gen_add_thmss pre_name = | 
| 174 | fold_map (add_thms_atts pre_name); | |
| 5907 | 175 | |
| 12235 
5fa04fc9b254
Further restructuring of theorem naming functions.
 berghofe parents: 
12138diff
changeset | 176 | fun gen_add_thms pre_name args = | 
| 18377 | 177 | apfst (map hd) o gen_add_thmss pre_name (map (apfst (apsnd single)) args); | 
| 12235 
5fa04fc9b254
Further restructuring of theorem naming functions.
 berghofe parents: 
12138diff
changeset | 178 | |
| 33700 | 179 | val add_thmss = gen_add_thmss (name_thms true true); | 
| 180 | val add_thms = gen_add_thms (name_thms true true); | |
| 27683 | 181 | val add_thm = yield_singleton add_thms; | 
| 5907 | 182 | |
| 183 | ||
| 26488 
b497e3187ec7
eliminated destructive/critical theorem database;
 wenzelm parents: 
26471diff
changeset | 184 | (* add_thms_dynamic *) | 
| 
b497e3187ec7
eliminated destructive/critical theorem database;
 wenzelm parents: 
26471diff
changeset | 185 | |
| 29579 | 186 | fun add_thms_dynamic (b, f) thy = thy | 
| 39557 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 wenzelm parents: 
39507diff
changeset | 187 | |> (Data.map o apfst) | 
| 42375 
774df7c59508
report Name_Space.declare/define, relatively to context;
 wenzelm parents: 
41696diff
changeset | 188 | (Facts.add_dynamic (Proof_Context.init_global thy) (Sign.naming_of thy) (b, f) #> snd); | 
| 26488 
b497e3187ec7
eliminated destructive/critical theorem database;
 wenzelm parents: 
26471diff
changeset | 189 | |
| 
b497e3187ec7
eliminated destructive/critical theorem database;
 wenzelm parents: 
26471diff
changeset | 190 | |
| 27728 | 191 | (* note_thmss *) | 
| 5907 | 192 | |
| 33167 | 193 | fun note_thmss kind = fold_map (fn ((b, more_atts), ths_atts) => fn thy => | 
| 12711 | 194 | let | 
| 28965 | 195 | val name = Sign.full_name thy b; | 
| 30190 | 196 | fun app (x, (ths, atts)) = Library.foldl_map (Thm.theory_attributes atts) (x, ths); | 
| 18418 
bf448d999b7e
re-arranged tuples (theory * 'a) to ('a * theory) in Pure
 haftmann parents: 
18377diff
changeset | 197 | val (thms, thy') = thy |> enter_thms | 
| 33700 | 198 | (name_thmss true) (name_thms false true) (apsnd flat o Library.foldl_map app) | 
| 33167 | 199 | (b, map (fn (ths, atts) => (ths, surround (Thm.kind kind) (atts @ more_atts))) ths_atts); | 
| 28076 | 200 | in ((name, thms), thy') end); | 
| 12711 | 201 | |
| 5280 | 202 | |
| 4022 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 wenzelm parents: 
4013diff
changeset | 203 | (* store axioms as theorems *) | 
| 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 wenzelm parents: 
4013diff
changeset | 204 | |
| 4853 | 205 | local | 
| 35985 
0bbf0d2348f9
moved Drule.forall_intr_frees to Thm.forall_intr_frees (in more_thm.ML, which is loaded before pure_thy.ML);
 wenzelm parents: 
35856diff
changeset | 206 | |
| 
0bbf0d2348f9
moved Drule.forall_intr_frees to Thm.forall_intr_frees (in more_thm.ML, which is loaded before pure_thy.ML);
 wenzelm parents: 
35856diff
changeset | 207 | fun no_read _ (_, t) = t; | 
| 
0bbf0d2348f9
moved Drule.forall_intr_frees to Thm.forall_intr_frees (in more_thm.ML, which is loaded before pure_thy.ML);
 wenzelm parents: 
35856diff
changeset | 208 | |
| 42375 
774df7c59508
report Name_Space.declare/define, relatively to context;
 wenzelm parents: 
41696diff
changeset | 209 | fun read ctxt (b, str) = | 
| 
774df7c59508
report Name_Space.declare/define, relatively to context;
 wenzelm parents: 
41696diff
changeset | 210 | Syntax.read_prop ctxt str handle ERROR msg => | 
| 42381 
309ec68442c6
added Binding.print convenience, which includes quote already;
 wenzelm parents: 
42379diff
changeset | 211 |     cat_error msg ("The error(s) above occurred in definition " ^ Binding.print b);
 | 
| 35985 
0bbf0d2348f9
moved Drule.forall_intr_frees to Thm.forall_intr_frees (in more_thm.ML, which is loaded before pure_thy.ML);
 wenzelm parents: 
35856diff
changeset | 212 | |
| 
0bbf0d2348f9
moved Drule.forall_intr_frees to Thm.forall_intr_frees (in more_thm.ML, which is loaded before pure_thy.ML);
 wenzelm parents: 
35856diff
changeset | 213 | fun add prep unchecked overloaded = fold_map (fn ((b, raw_prop), atts) => fn thy => | 
| 
0bbf0d2348f9
moved Drule.forall_intr_frees to Thm.forall_intr_frees (in more_thm.ML, which is loaded before pure_thy.ML);
 wenzelm parents: 
35856diff
changeset | 214 | let | 
| 42375 
774df7c59508
report Name_Space.declare/define, relatively to context;
 wenzelm parents: 
41696diff
changeset | 215 | val ctxt = Syntax.init_pretty_global thy; | 
| 
774df7c59508
report Name_Space.declare/define, relatively to context;
 wenzelm parents: 
41696diff
changeset | 216 | val prop = prep ctxt (b, raw_prop); | 
| 
774df7c59508
report Name_Space.declare/define, relatively to context;
 wenzelm parents: 
41696diff
changeset | 217 | val ((_, def), thy') = Thm.add_def ctxt unchecked overloaded (b, prop) thy; | 
| 35985 
0bbf0d2348f9
moved Drule.forall_intr_frees to Thm.forall_intr_frees (in more_thm.ML, which is loaded before pure_thy.ML);
 wenzelm parents: 
35856diff
changeset | 218 | val thm = def | 
| 
0bbf0d2348f9
moved Drule.forall_intr_frees to Thm.forall_intr_frees (in more_thm.ML, which is loaded before pure_thy.ML);
 wenzelm parents: 
35856diff
changeset | 219 | |> Thm.forall_intr_frees | 
| 
0bbf0d2348f9
moved Drule.forall_intr_frees to Thm.forall_intr_frees (in more_thm.ML, which is loaded before pure_thy.ML);
 wenzelm parents: 
35856diff
changeset | 220 | |> Thm.forall_elim_vars 0 | 
| 
0bbf0d2348f9
moved Drule.forall_intr_frees to Thm.forall_intr_frees (in more_thm.ML, which is loaded before pure_thy.ML);
 wenzelm parents: 
35856diff
changeset | 221 | |> Thm.varifyT_global; | 
| 
0bbf0d2348f9
moved Drule.forall_intr_frees to Thm.forall_intr_frees (in more_thm.ML, which is loaded before pure_thy.ML);
 wenzelm parents: 
35856diff
changeset | 222 | in yield_singleton (gen_add_thms (K I)) ((b, thm), atts) thy' end); | 
| 
0bbf0d2348f9
moved Drule.forall_intr_frees to Thm.forall_intr_frees (in more_thm.ML, which is loaded before pure_thy.ML);
 wenzelm parents: 
35856diff
changeset | 223 | |
| 4853 | 224 | in | 
| 35985 
0bbf0d2348f9
moved Drule.forall_intr_frees to Thm.forall_intr_frees (in more_thm.ML, which is loaded before pure_thy.ML);
 wenzelm parents: 
35856diff
changeset | 225 | |
| 
0bbf0d2348f9
moved Drule.forall_intr_frees to Thm.forall_intr_frees (in more_thm.ML, which is loaded before pure_thy.ML);
 wenzelm parents: 
35856diff
changeset | 226 | val add_defs = add no_read false; | 
| 
0bbf0d2348f9
moved Drule.forall_intr_frees to Thm.forall_intr_frees (in more_thm.ML, which is loaded before pure_thy.ML);
 wenzelm parents: 
35856diff
changeset | 227 | val add_defs_unchecked = add no_read true; | 
| 
0bbf0d2348f9
moved Drule.forall_intr_frees to Thm.forall_intr_frees (in more_thm.ML, which is loaded before pure_thy.ML);
 wenzelm parents: 
35856diff
changeset | 228 | val add_defs_cmd = add read false; | 
| 
0bbf0d2348f9
moved Drule.forall_intr_frees to Thm.forall_intr_frees (in more_thm.ML, which is loaded before pure_thy.ML);
 wenzelm parents: 
35856diff
changeset | 229 | val add_defs_unchecked_cmd = add read true; | 
| 
0bbf0d2348f9
moved Drule.forall_intr_frees to Thm.forall_intr_frees (in more_thm.ML, which is loaded before pure_thy.ML);
 wenzelm parents: 
35856diff
changeset | 230 | |
| 4853 | 231 | end; | 
| 4022 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 wenzelm parents: 
4013diff
changeset | 232 | |
| 3987 | 233 | end; |