| author | Fabian Huch <huch@in.tum.de> | 
| Fri, 03 Nov 2023 10:21:21 +0100 | |
| changeset 78887 | 6996a20a1b7c | 
| parent 77979 | a12c48fbf10f | 
| child 79113 | 5109e4b2a292 | 
| 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 | 
| 74261 
d28a51dd9da6
export other entities, e.g. relevant for formal document output;
 wenzelm parents: 
74232diff
changeset | 10 | val fact_space: theory -> Name_Space.T | 
| 56003 | 11 | val check_fact: theory -> xstring * Position.T -> string | 
| 26666 | 12 | val intern_fact: theory -> xstring -> string | 
| 26693 | 13 | val defined_fact: theory -> string -> bool | 
| 57887 | 14 | val alias_fact: binding -> string -> theory -> theory | 
| 27198 | 15 | val hide_fact: bool -> string -> theory -> theory | 
| 70574 | 16 | val dest_thms: bool -> theory list -> theory -> (Thm_Name.T * thm) list | 
| 72098 
8c547eac8381
more robust treatment of thm_names, with strict check after all theories are loaded;
 wenzelm parents: 
72055diff
changeset | 17 | val get_thm_names: theory -> Thm_Name.T Inttab.table | 
| 70904 | 18 | val dest_thm_names: theory -> (Proofterm.thm_id * Thm_Name.T) list | 
| 70574 | 19 | val lookup_thm_id: theory -> Proofterm.thm_id -> Thm_Name.T option | 
| 70594 | 20 | val lookup_thm: theory -> thm -> (Proofterm.thm_id * Thm_Name.T) option | 
| 26344 
04dacc6809b6
simplified get_thm(s): back to plain name argument;
 wenzelm parents: 
26336diff
changeset | 21 | val get_thms: theory -> xstring -> thm list | 
| 
04dacc6809b6
simplified get_thm(s): back to plain name argument;
 wenzelm parents: 
26336diff
changeset | 22 | val get_thm: theory -> xstring -> thm | 
| 61054 | 23 | val transfer_theories: theory -> thm -> thm | 
| 56161 | 24 | val all_thms_of: theory -> bool -> (string * thm) list | 
| 70574 | 25 | val get_thm_name: theory -> Thm_Name.T * Position.T -> thm | 
| 21567 | 26 |   val burrow_fact: ('a list -> 'b list) -> ('a list * 'c) list -> ('b list * 'c) list
 | 
| 21580 | 27 |   val burrow_facts: ('a list -> 'b list) ->
 | 
| 28 |     ('c * ('a list * 'd) list) list -> ('c * ('b list * 'd) list) list
 | |
| 70550 | 29 | val name_multi: string * Position.T -> thm list -> ((string * Position.T) * thm) list | 
| 70430 | 30 | type name_flags | 
| 31 | val unnamed: name_flags | |
| 70427 | 32 | val official1: name_flags | 
| 33 | val official2: name_flags | |
| 34 | val unofficial1: name_flags | |
| 35 | val unofficial2: name_flags | |
| 70494 | 36 | val name_thm: name_flags -> string * Position.T -> thm -> thm | 
| 37 | val name_thms: name_flags -> string * Position.T -> thm list -> thm list | |
| 68661 | 38 | val check_thms_lazy: thm list lazy -> thm list lazy | 
| 67671 
857da80611ab
support for lazy notes in global/local context and Element.Lazy_Notes: name binding and fact without attributes;
 wenzelm parents: 
67663diff
changeset | 39 | val add_thms_lazy: string -> (binding * thm list lazy) -> theory -> theory | 
| 29579 | 40 | val store_thm: binding * thm -> theory -> thm * theory | 
| 41 | val store_thm_open: binding * thm -> theory -> thm * theory | |
| 42 | val add_thms: ((binding * thm) * attribute list) list -> theory -> thm list * theory | |
| 43 | val add_thm: (binding * thm) * attribute list -> theory -> thm * theory | |
| 44 | val add_thmss: ((binding * thm list) * attribute list) list -> theory -> thm list list * theory | |
| 57929 
c5063c033a5a
tuned signature -- proper Local_Theory.add_thms_dynamic;
 wenzelm parents: 
57887diff
changeset | 45 | val add_thms_dynamic': Context.generic -> binding * (Context.generic -> thm list) -> | 
| 
c5063c033a5a
tuned signature -- proper Local_Theory.add_thms_dynamic;
 wenzelm parents: 
57887diff
changeset | 46 | theory -> string * theory | 
| 29579 | 47 | val add_thms_dynamic: binding * (Context.generic -> thm list) -> theory -> theory | 
| 67713 | 48 | val note_thms: string -> Thm.binding * (thm list * attribute list) list -> theory -> | 
| 49 | (string * thm list) * theory | |
| 50 | val note_thmss: string -> (Thm.binding * (thm list * attribute list) list) list -> theory -> | |
| 51 | (string * thm list) list * theory | |
| 29579 | 52 | val add_defs: bool -> ((binding * term) * attribute list) list -> | 
| 18377 | 53 | theory -> thm list * theory | 
| 29579 | 54 | val add_defs_unchecked: bool -> ((binding * term) * attribute list) list -> | 
| 55 | theory -> thm list * theory | |
| 3987 | 56 | end; | 
| 57 | ||
| 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 | 58 | structure Global_Theory: GLOBAL_THEORY = | 
| 3987 | 59 | struct | 
| 60 | ||
| 27198 | 61 | (** 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 | 62 | |
| 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 | 63 | structure Data = Theory_Data | 
| 24713 | 64 | ( | 
| 70574 | 65 | type T = Facts.T * Thm_Name.T Inttab.table lazy option; | 
| 70552 
8d7a531a6b58
maintain thm_name vs. derivation_id for global facts;
 wenzelm parents: 
70550diff
changeset | 66 | val empty: T = (Facts.empty, NONE); | 
| 
8d7a531a6b58
maintain thm_name vs. derivation_id for global facts;
 wenzelm parents: 
70550diff
changeset | 67 | fun merge ((facts1, _), (facts2, _)) = (Facts.merge (facts1, facts2), NONE); | 
| 24713 | 68 | ); | 
| 3987 | 69 | |
| 70552 
8d7a531a6b58
maintain thm_name vs. derivation_id for global facts;
 wenzelm parents: 
70550diff
changeset | 70 | |
| 
8d7a531a6b58
maintain thm_name vs. derivation_id for global facts;
 wenzelm parents: 
70550diff
changeset | 71 | (* facts *) | 
| 
8d7a531a6b58
maintain thm_name vs. derivation_id for global facts;
 wenzelm parents: 
70550diff
changeset | 72 | |
| 
8d7a531a6b58
maintain thm_name vs. derivation_id for global facts;
 wenzelm parents: 
70550diff
changeset | 73 | val facts_of = #1 o Data.get; | 
| 
8d7a531a6b58
maintain thm_name vs. derivation_id for global facts;
 wenzelm parents: 
70550diff
changeset | 74 | val map_facts = Data.map o apfst; | 
| 26666 | 75 | |
| 74261 
d28a51dd9da6
export other entities, e.g. relevant for formal document output;
 wenzelm parents: 
74232diff
changeset | 76 | val fact_space = Facts.space_of o facts_of; | 
| 56003 | 77 | fun check_fact thy = Facts.check (Context.Theory thy) (facts_of thy); | 
| 26666 | 78 | val intern_fact = Facts.intern o facts_of; | 
| 26693 | 79 | val defined_fact = Facts.defined o facts_of; | 
| 16023 
66561f6814bd
added string_of_thmref, selections, fact_index_of, valid_thms;
 wenzelm parents: 
15975diff
changeset | 80 | |
| 57887 | 81 | fun alias_fact binding name thy = | 
| 77979 
a12c48fbf10f
back to more elementary concept of aliases as adhoc change of accesses, but now with "suppress" information (see also 31ea5c1f874d);
 wenzelm parents: 
77970diff
changeset | 82 | map_facts (Facts.alias (Sign.naming_of thy) binding name) thy; | 
| 57887 | 83 | |
| 70544 | 84 | fun hide_fact fully name = map_facts (Facts.hide fully name); | 
| 6367 | 85 | |
| 70552 
8d7a531a6b58
maintain thm_name vs. derivation_id for global facts;
 wenzelm parents: 
70550diff
changeset | 86 | fun dest_thms verbose prev_thys thy = | 
| 
8d7a531a6b58
maintain thm_name vs. derivation_id for global facts;
 wenzelm parents: 
70550diff
changeset | 87 | Facts.dest_static verbose (map facts_of prev_thys) (facts_of thy) | 
| 70574 | 88 | |> maps (uncurry Thm_Name.make_list); | 
| 70552 
8d7a531a6b58
maintain thm_name vs. derivation_id for global facts;
 wenzelm parents: 
70550diff
changeset | 89 | |
| 
8d7a531a6b58
maintain thm_name vs. derivation_id for global facts;
 wenzelm parents: 
70550diff
changeset | 90 | |
| 
8d7a531a6b58
maintain thm_name vs. derivation_id for global facts;
 wenzelm parents: 
70550diff
changeset | 91 | (* thm_name vs. derivation_id *) | 
| 
8d7a531a6b58
maintain thm_name vs. derivation_id for global facts;
 wenzelm parents: 
70550diff
changeset | 92 | |
| 
8d7a531a6b58
maintain thm_name vs. derivation_id for global facts;
 wenzelm parents: 
70550diff
changeset | 93 | val thm_names_of = #2 o Data.get; | 
| 
8d7a531a6b58
maintain thm_name vs. derivation_id for global facts;
 wenzelm parents: 
70550diff
changeset | 94 | val map_thm_names = Data.map o apsnd; | 
| 
8d7a531a6b58
maintain thm_name vs. derivation_id for global facts;
 wenzelm parents: 
70550diff
changeset | 95 | |
| 
8d7a531a6b58
maintain thm_name vs. derivation_id for global facts;
 wenzelm parents: 
70550diff
changeset | 96 | fun make_thm_names thy = | 
| 
8d7a531a6b58
maintain thm_name vs. derivation_id for global facts;
 wenzelm parents: 
70550diff
changeset | 97 | (dest_thms true (Theory.parents_of thy) thy, Inttab.empty) | 
| 
8d7a531a6b58
maintain thm_name vs. derivation_id for global facts;
 wenzelm parents: 
70550diff
changeset | 98 | |-> fold (fn (thm_name, thm) => fn thm_names => | 
| 
8d7a531a6b58
maintain thm_name vs. derivation_id for global facts;
 wenzelm parents: 
70550diff
changeset | 99 | (case Thm.derivation_id (Thm.transfer thy thm) of | 
| 
8d7a531a6b58
maintain thm_name vs. derivation_id for global facts;
 wenzelm parents: 
70550diff
changeset | 100 | NONE => thm_names | 
| 
8d7a531a6b58
maintain thm_name vs. derivation_id for global facts;
 wenzelm parents: 
70550diff
changeset | 101 |     | SOME {serial, theory_name} =>
 | 
| 
8d7a531a6b58
maintain thm_name vs. derivation_id for global facts;
 wenzelm parents: 
70550diff
changeset | 102 | if Context.theory_long_name thy <> theory_name then | 
| 
8d7a531a6b58
maintain thm_name vs. derivation_id for global facts;
 wenzelm parents: 
70550diff
changeset | 103 |           raise THM ("Bad theory name for derivation", 0, [thm])
 | 
| 
8d7a531a6b58
maintain thm_name vs. derivation_id for global facts;
 wenzelm parents: 
70550diff
changeset | 104 | else | 
| 
8d7a531a6b58
maintain thm_name vs. derivation_id for global facts;
 wenzelm parents: 
70550diff
changeset | 105 | (case Inttab.lookup thm_names serial of | 
| 
8d7a531a6b58
maintain thm_name vs. derivation_id for global facts;
 wenzelm parents: 
70550diff
changeset | 106 | SOME thm_name' => | 
| 
8d7a531a6b58
maintain thm_name vs. derivation_id for global facts;
 wenzelm parents: 
70550diff
changeset | 107 |               raise THM ("Duplicate use of derivation identity for " ^
 | 
| 70574 | 108 | Thm_Name.print thm_name ^ " vs. " ^ | 
| 109 | Thm_Name.print thm_name', 0, [thm]) | |
| 70552 
8d7a531a6b58
maintain thm_name vs. derivation_id for global facts;
 wenzelm parents: 
70550diff
changeset | 110 | | NONE => Inttab.update (serial, thm_name) thm_names))); | 
| 
8d7a531a6b58
maintain thm_name vs. derivation_id for global facts;
 wenzelm parents: 
70550diff
changeset | 111 | |
| 72098 
8c547eac8381
more robust treatment of thm_names, with strict check after all theories are loaded;
 wenzelm parents: 
72055diff
changeset | 112 | fun lazy_thm_names thy = | 
| 70552 
8d7a531a6b58
maintain thm_name vs. derivation_id for global facts;
 wenzelm parents: 
70550diff
changeset | 113 | (case thm_names_of thy of | 
| 70555 
c1fde53e5e82
clarified lookup operations: more scalable for multiple retrieval;
 wenzelm parents: 
70552diff
changeset | 114 | NONE => Lazy.lazy (fn () => make_thm_names thy) | 
| 
c1fde53e5e82
clarified lookup operations: more scalable for multiple retrieval;
 wenzelm parents: 
70552diff
changeset | 115 | | SOME lazy_tab => lazy_tab); | 
| 70552 
8d7a531a6b58
maintain thm_name vs. derivation_id for global facts;
 wenzelm parents: 
70550diff
changeset | 116 | |
| 72098 
8c547eac8381
more robust treatment of thm_names, with strict check after all theories are loaded;
 wenzelm parents: 
72055diff
changeset | 117 | val get_thm_names = Lazy.force o lazy_thm_names; | 
| 
8c547eac8381
more robust treatment of thm_names, with strict check after all theories are loaded;
 wenzelm parents: 
72055diff
changeset | 118 | |
| 70904 | 119 | fun dest_thm_names thy = | 
| 120 | let | |
| 121 | val theory_name = Context.theory_long_name thy; | |
| 122 | fun thm_id i = Proofterm.make_thm_id (i, theory_name); | |
| 72098 
8c547eac8381
more robust treatment of thm_names, with strict check after all theories are loaded;
 wenzelm parents: 
72055diff
changeset | 123 | in Inttab.fold_rev (fn (i, thm_name) => cons (thm_id i, thm_name)) (get_thm_names thy) [] end; | 
| 70552 
8d7a531a6b58
maintain thm_name vs. derivation_id for global facts;
 wenzelm parents: 
70550diff
changeset | 124 | |
| 70555 
c1fde53e5e82
clarified lookup operations: more scalable for multiple retrieval;
 wenzelm parents: 
70552diff
changeset | 125 | fun lookup_thm_id thy = | 
| 
c1fde53e5e82
clarified lookup operations: more scalable for multiple retrieval;
 wenzelm parents: 
70552diff
changeset | 126 | let | 
| 
c1fde53e5e82
clarified lookup operations: more scalable for multiple retrieval;
 wenzelm parents: 
70552diff
changeset | 127 | val theories = | 
| 74232 | 128 | Symtab.build (Theory.nodes_of thy |> fold (fn thy' => | 
| 129 | Symtab.update (Context.theory_long_name thy', lazy_thm_names thy'))); | |
| 70555 
c1fde53e5e82
clarified lookup operations: more scalable for multiple retrieval;
 wenzelm parents: 
70552diff
changeset | 130 | fun lookup (thm_id: Proofterm.thm_id) = | 
| 
c1fde53e5e82
clarified lookup operations: more scalable for multiple retrieval;
 wenzelm parents: 
70552diff
changeset | 131 | (case Symtab.lookup theories (#theory_name thm_id) of | 
| 
c1fde53e5e82
clarified lookup operations: more scalable for multiple retrieval;
 wenzelm parents: 
70552diff
changeset | 132 | NONE => NONE | 
| 
c1fde53e5e82
clarified lookup operations: more scalable for multiple retrieval;
 wenzelm parents: 
70552diff
changeset | 133 | | SOME lazy_tab => Inttab.lookup (Lazy.force lazy_tab) (#serial thm_id)); | 
| 
c1fde53e5e82
clarified lookup operations: more scalable for multiple retrieval;
 wenzelm parents: 
70552diff
changeset | 134 | in lookup end; | 
| 
c1fde53e5e82
clarified lookup operations: more scalable for multiple retrieval;
 wenzelm parents: 
70552diff
changeset | 135 | |
| 
c1fde53e5e82
clarified lookup operations: more scalable for multiple retrieval;
 wenzelm parents: 
70552diff
changeset | 136 | fun lookup_thm thy = | 
| 70594 | 137 | let val lookup = lookup_thm_id thy in | 
| 138 | fn thm => | |
| 139 | (case Thm.derivation_id thm of | |
| 140 | NONE => NONE | |
| 141 | | SOME thm_id => | |
| 142 | (case lookup thm_id of | |
| 143 | NONE => NONE | |
| 144 | | SOME thm_name => SOME (thm_id, thm_name))) | |
| 145 | end; | |
| 70552 
8d7a531a6b58
maintain thm_name vs. derivation_id for global facts;
 wenzelm parents: 
70550diff
changeset | 146 | |
| 
8d7a531a6b58
maintain thm_name vs. derivation_id for global facts;
 wenzelm parents: 
70550diff
changeset | 147 | val _ = | 
| 72098 
8c547eac8381
more robust treatment of thm_names, with strict check after all theories are loaded;
 wenzelm parents: 
72055diff
changeset | 148 | Theory.setup | 
| 
8c547eac8381
more robust treatment of thm_names, with strict check after all theories are loaded;
 wenzelm parents: 
72055diff
changeset | 149 | (Theory.at_begin (fn thy => | 
| 
8c547eac8381
more robust treatment of thm_names, with strict check after all theories are loaded;
 wenzelm parents: 
72055diff
changeset | 150 | if is_none (thm_names_of thy) then NONE | 
| 
8c547eac8381
more robust treatment of thm_names, with strict check after all theories are loaded;
 wenzelm parents: 
72055diff
changeset | 151 | else SOME (map_thm_names (K NONE) thy)) #> | 
| 
8c547eac8381
more robust treatment of thm_names, with strict check after all theories are loaded;
 wenzelm parents: 
72055diff
changeset | 152 | Theory.at_end (fn thy => | 
| 
8c547eac8381
more robust treatment of thm_names, with strict check after all theories are loaded;
 wenzelm parents: 
72055diff
changeset | 153 | if is_some (thm_names_of thy) then NONE | 
| 
8c547eac8381
more robust treatment of thm_names, with strict check after all theories are loaded;
 wenzelm parents: 
72055diff
changeset | 154 | else | 
| 
8c547eac8381
more robust treatment of thm_names, with strict check after all theories are loaded;
 wenzelm parents: 
72055diff
changeset | 155 | let | 
| 
8c547eac8381
more robust treatment of thm_names, with strict check after all theories are loaded;
 wenzelm parents: 
72055diff
changeset | 156 | val lazy_tab = | 
| 
8c547eac8381
more robust treatment of thm_names, with strict check after all theories are loaded;
 wenzelm parents: 
72055diff
changeset | 157 | if Future.proofs_enabled 1 | 
| 
8c547eac8381
more robust treatment of thm_names, with strict check after all theories are loaded;
 wenzelm parents: 
72055diff
changeset | 158 | then Lazy.lazy (fn () => make_thm_names thy) | 
| 
8c547eac8381
more robust treatment of thm_names, with strict check after all theories are loaded;
 wenzelm parents: 
72055diff
changeset | 159 | else Lazy.value (make_thm_names thy); | 
| 
8c547eac8381
more robust treatment of thm_names, with strict check after all theories are loaded;
 wenzelm parents: 
72055diff
changeset | 160 | in SOME (map_thm_names (K (SOME lazy_tab)) thy) end)); | 
| 70552 
8d7a531a6b58
maintain thm_name vs. derivation_id for global facts;
 wenzelm parents: 
70550diff
changeset | 161 | |
| 3987 | 162 | |
| 56140 
ed92ce2ac88e
just one cumulative Proof_Context.facts, with uniform retrieval (including PIDE markup, completion etc.);
 wenzelm parents: 
56003diff
changeset | 163 | (* retrieve theorems *) | 
| 27198 | 164 | |
| 56140 
ed92ce2ac88e
just one cumulative Proof_Context.facts, with uniform retrieval (including PIDE markup, completion etc.);
 wenzelm parents: 
56003diff
changeset | 165 | fun get_thms thy xname = | 
| 57942 
e5bec882fdd0
more informative Token.Fact: retain name of dynamic fact (without selection);
 wenzelm parents: 
57929diff
changeset | 166 | #thms (Facts.retrieve (Context.Theory thy) (facts_of thy) (xname, Position.none)); | 
| 26344 
04dacc6809b6
simplified get_thm(s): back to plain name argument;
 wenzelm parents: 
26336diff
changeset | 167 | |
| 56140 
ed92ce2ac88e
just one cumulative Proof_Context.facts, with uniform retrieval (including PIDE markup, completion etc.);
 wenzelm parents: 
56003diff
changeset | 168 | fun get_thm thy xname = | 
| 
ed92ce2ac88e
just one cumulative Proof_Context.facts, with uniform retrieval (including PIDE markup, completion etc.);
 wenzelm parents: 
56003diff
changeset | 169 | Facts.the_single (xname, Position.none) (get_thms thy xname); | 
| 4783 | 170 | |
| 61054 | 171 | fun transfer_theories thy = | 
| 172 | let | |
| 173 | val theories = | |
| 74232 | 174 | Symtab.build (Theory.nodes_of thy |> fold (fn thy' => | 
| 77890 | 175 | Symtab.update (Context.theory_long_name thy', thy'))); | 
| 61054 | 176 | fun transfer th = | 
| 77890 | 177 | Thm.transfer (the_default thy (Symtab.lookup theories (Thm.theory_long_name th))) th; | 
| 61054 | 178 | in transfer end; | 
| 179 | ||
| 56161 | 180 | fun all_thms_of thy verbose = | 
| 181 | let | |
| 61054 | 182 | val transfer = transfer_theories thy; | 
| 56161 | 183 | val facts = facts_of thy; | 
| 184 | fun add (name, ths) = | |
| 185 | if not verbose andalso Facts.is_concealed facts name then I | |
| 61054 | 186 | else append (map (`(Thm.get_name_hint) o transfer) ths); | 
| 56161 | 187 | in Facts.fold_static add facts [] end; | 
| 16336 | 188 | |
| 70574 | 189 | fun get_thm_name thy ((name, i), pos) = | 
| 190 | let | |
| 191 | val facts = facts_of thy; | |
| 192 | fun print_name () = | |
| 193 | Facts.markup_extern (Proof_Context.init_global thy) facts name |-> Markup.markup; | |
| 194 | in | |
| 195 | (case (Facts.lookup (Context.Theory thy) facts name, i) of | |
| 196 |       (NONE, _) => error ("Undefined fact " ^ quote (print_name ()) ^ Position.here pos)
 | |
| 197 |     | (SOME {thms = [thm], ...}, 0) => thm
 | |
| 198 |     | (SOME {thms, ...}, 0) => Facts.err_single (print_name (), pos) thms
 | |
| 199 |     | (SOME {thms, ...}, _) =>
 | |
| 200 | if i > 0 andalso i <= length thms then nth thms (i - 1) | |
| 201 | else Facts.err_selection (print_name (), pos) i thms) | |
| 70894 | 202 | |> Thm.transfer thy | 
| 70574 | 203 | end; | 
| 204 | ||
| 4022 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 wenzelm parents: 
4013diff
changeset | 205 | |
| 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 wenzelm parents: 
4013diff
changeset | 206 | |
| 26488 
b497e3187ec7
eliminated destructive/critical theorem database;
 wenzelm parents: 
26471diff
changeset | 207 | (** store theorems **) | 
| 3987 | 208 | |
| 21580 | 209 | (* fact specifications *) | 
| 210 | ||
| 211 | fun burrow_fact f = split_list #>> burrow f #> op ~~; | |
| 212 | fun burrow_facts f = split_list ##> burrow (burrow_fact f) #> op ~~; | |
| 213 | ||
| 214 | ||
| 70430 | 215 | (* name theorems *) | 
| 4853 | 216 | |
| 70430 | 217 | abstype name_flags = No_Name_Flags | Name_Flags of {pre: bool, official: bool}
 | 
| 218 | with | |
| 12235 
5fa04fc9b254
Further restructuring of theorem naming functions.
 berghofe parents: 
12138diff
changeset | 219 | |
| 70430 | 220 | val unnamed = No_Name_Flags; | 
| 221 | val official1 = Name_Flags {pre = true, official = true};
 | |
| 222 | val official2 = Name_Flags {pre = false, official = true};
 | |
| 223 | val unofficial1 = Name_Flags {pre = true, official = false};
 | |
| 224 | val unofficial2 = Name_Flags {pre = false, official = false};
 | |
| 70427 | 225 | |
| 70494 | 226 | fun name_thm name_flags (name, pos) = | 
| 70459 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 wenzelm parents: 
70431diff
changeset | 227 | Thm.solve_constraints #> (fn thm => | 
| 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 wenzelm parents: 
70431diff
changeset | 228 | (case name_flags of | 
| 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 wenzelm parents: 
70431diff
changeset | 229 | No_Name_Flags => thm | 
| 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 wenzelm parents: 
70431diff
changeset | 230 |     | Name_Flags {pre, official} =>
 | 
| 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 wenzelm parents: 
70431diff
changeset | 231 | thm | 
| 70543 
33749040b6f8
clarified derivation_name vs. raw_derivation_name;
 wenzelm parents: 
70494diff
changeset | 232 | |> (official andalso (not pre orelse Thm.raw_derivation_name thm = "")) ? | 
| 70494 | 233 | Thm.name_derivation (name, pos) | 
| 70459 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 wenzelm parents: 
70431diff
changeset | 234 | |> (name <> "" andalso (not pre orelse not (Thm.has_name_hint thm))) ? | 
| 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 wenzelm parents: 
70431diff
changeset | 235 | Thm.put_name_hint name)); | 
| 12872 
0855c3ab2047
Theorems are only "pre-named" if the do not already have names.
 berghofe parents: 
12711diff
changeset | 236 | |
| 70430 | 237 | end; | 
| 70427 | 238 | |
| 70550 | 239 | fun name_multi (name, pos) = | 
| 70574 | 240 | Thm_Name.make_list name #> (map o apfst) (fn thm_name => (Thm_Name.flatten thm_name, pos)); | 
| 70427 | 241 | |
| 70550 | 242 | fun name_thms name_flags name_pos = | 
| 243 | name_multi name_pos #> map (uncurry (name_thm name_flags)); | |
| 12235 
5fa04fc9b254
Further restructuring of theorem naming functions.
 berghofe parents: 
12138diff
changeset | 244 | |
| 4853 | 245 | |
| 70424 | 246 | (* apply theorems and attributes *) | 
| 4853 | 247 | |
| 67671 
857da80611ab
support for lazy notes in global/local context and Element.Lazy_Notes: name binding and fact without attributes;
 wenzelm parents: 
67663diff
changeset | 248 | fun register_proofs thms thy = (thms, Thm.register_proofs (Lazy.value thms) thy); | 
| 49010 | 249 | |
| 70494 | 250 | fun bind_name thy b = (Sign.full_name thy b, Binding.default_pos_of b); | 
| 251 | ||
| 68244 
e0cd57aeb60c
more checks for global facts: disallow undeclared frees (as in Export_Theory.export_fact);
 wenzelm parents: 
67779diff
changeset | 252 | fun add_facts (b, fact) thy = | 
| 
e0cd57aeb60c
more checks for global facts: disallow undeclared frees (as in Export_Theory.export_fact);
 wenzelm parents: 
67779diff
changeset | 253 | let | 
| 70494 | 254 | val (full_name, pos) = bind_name thy b; | 
| 68540 | 255 | fun check fact = | 
| 256 | fact |> map_index (fn (i, thm) => | |
| 257 | let | |
| 258 | fun err msg = | |
| 259 |             error ("Malformed global fact " ^
 | |
| 260 | quote (full_name ^ | |
| 261 |                 (if length fact = 1 then "" else "(" ^ string_of_int (i + 1) ^ ")")) ^
 | |
| 262 | Position.here pos ^ "\n" ^ msg); | |
| 263 | val prop = Thm.plain_prop_of thm | |
| 264 | handle THM _ => | |
| 265 | thm | |
| 266 | |> Thm.check_hyps (Context.Theory thy) | |
| 267 | |> Thm.full_prop_of; | |
| 268 | in | |
| 269 | ignore (Logic.unvarify_global (Term_Subst.zero_var_indexes prop)) | |
| 270 | handle TYPE (msg, _, _) => err msg | |
| 271 | | TERM (msg, _) => err msg | |
| 272 | | ERROR msg => err msg | |
| 273 | end); | |
| 274 | val arg = (b, Lazy.map_finished (tap check) fact); | |
| 68244 
e0cd57aeb60c
more checks for global facts: disallow undeclared frees (as in Export_Theory.export_fact);
 wenzelm parents: 
67779diff
changeset | 275 | in | 
| 70544 | 276 |     thy |> map_facts (Facts.add_static (Context.Theory thy) {strict = true, index = false} arg #> #2)
 | 
| 68244 
e0cd57aeb60c
more checks for global facts: disallow undeclared frees (as in Export_Theory.export_fact);
 wenzelm parents: 
67779diff
changeset | 277 | end; | 
| 67671 
857da80611ab
support for lazy notes in global/local context and Element.Lazy_Notes: name binding and fact without attributes;
 wenzelm parents: 
67663diff
changeset | 278 | |
| 68661 | 279 | fun check_thms_lazy (thms: thm list lazy) = | 
| 68701 
be936cf061ab
more robust: do not defer potentially slow/big lazy facts to the very end;
 wenzelm parents: 
68699diff
changeset | 280 | if Proofterm.proofs_enabled () orelse Options.default_bool "strict_facts" | 
| 
be936cf061ab
more robust: do not defer potentially slow/big lazy facts to the very end;
 wenzelm parents: 
68699diff
changeset | 281 | then Lazy.force_value thms else thms; | 
| 68661 | 282 | |
| 67671 
857da80611ab
support for lazy notes in global/local context and Element.Lazy_Notes: name binding and fact without attributes;
 wenzelm parents: 
67663diff
changeset | 283 | fun add_thms_lazy kind (b, thms) thy = | 
| 68661 | 284 | if Binding.is_empty b then Thm.register_proofs (check_thms_lazy thms) thy | 
| 30211 | 285 | else | 
| 286 | let | |
| 70494 | 287 | val name_pos = bind_name thy b; | 
| 68661 | 288 | val thms' = | 
| 289 | check_thms_lazy thms | |
| 70494 | 290 | |> Lazy.map_finished (name_thms official1 name_pos #> map (Thm.kind_rule kind)); | 
| 67671 
857da80611ab
support for lazy notes in global/local context and Element.Lazy_Notes: name binding and fact without attributes;
 wenzelm parents: 
67663diff
changeset | 291 | in thy |> Thm.register_proofs thms' |> add_facts (b, thms') end; | 
| 
857da80611ab
support for lazy notes in global/local context and Element.Lazy_Notes: name binding and fact without attributes;
 wenzelm parents: 
67663diff
changeset | 292 | |
| 70424 | 293 | val app_facts = | 
| 70552 
8d7a531a6b58
maintain thm_name vs. derivation_id for global facts;
 wenzelm parents: 
70550diff
changeset | 294 | apfst flat oo fold_map (fn (thms, atts) => fn thy => | 
| 
8d7a531a6b58
maintain thm_name vs. derivation_id for global facts;
 wenzelm parents: 
70550diff
changeset | 295 | fold_map (Thm.theory_attributes atts) (map (Thm.transfer thy) thms) thy); | 
| 70424 | 296 | |
| 70430 | 297 | fun apply_facts name_flags1 name_flags2 (b, facts) thy = | 
| 71018 
d32ed8927a42
more robust expose_proofs corresponding to register_proofs/consolidate_theory;
 wenzelm parents: 
71006diff
changeset | 298 | if Binding.is_empty b then app_facts facts thy |-> register_proofs | 
| 67671 
857da80611ab
support for lazy notes in global/local context and Element.Lazy_Notes: name binding and fact without attributes;
 wenzelm parents: 
67663diff
changeset | 299 | else | 
| 
857da80611ab
support for lazy notes in global/local context and Element.Lazy_Notes: name binding and fact without attributes;
 wenzelm parents: 
67663diff
changeset | 300 | let | 
| 70985 
2866fee241ee
proper export of unnamed proof boxes for unnamed toplevel declarations, e.g. rulify/defn rules in theory IFOL and HOL;
 wenzelm parents: 
70904diff
changeset | 301 | val name_pos = bind_name thy b; | 
| 70424 | 302 | val (thms', thy') = thy | 
| 70494 | 303 | |> app_facts (map (apfst (name_thms name_flags1 name_pos)) facts) | 
| 304 | |>> name_thms name_flags2 name_pos |-> register_proofs; | |
| 67671 
857da80611ab
support for lazy notes in global/local context and Element.Lazy_Notes: name binding and fact without attributes;
 wenzelm parents: 
67663diff
changeset | 305 | val thy'' = thy' |> add_facts (b, Lazy.value thms'); | 
| 67663 | 306 | in (map (Thm.transfer thy'') thms', thy'') end; | 
| 26488 
b497e3187ec7
eliminated destructive/critical theorem database;
 wenzelm parents: 
26471diff
changeset | 307 | |
| 
b497e3187ec7
eliminated destructive/critical theorem database;
 wenzelm parents: 
26471diff
changeset | 308 | |
| 67715 | 309 | (* store_thm *) | 
| 26488 
b497e3187ec7
eliminated destructive/critical theorem database;
 wenzelm parents: 
26471diff
changeset | 310 | |
| 67715 | 311 | fun store_thm (b, th) = | 
| 70430 | 312 | apply_facts official1 official2 (b, [([th], [])]) #>> the_single; | 
| 26488 
b497e3187ec7
eliminated destructive/critical theorem database;
 wenzelm parents: 
26471diff
changeset | 313 | |
| 29579 | 314 | fun store_thm_open (b, th) = | 
| 70430 | 315 | apply_facts unofficial1 unofficial2 (b, [([th], [])]) #>> the_single; | 
| 3987 | 316 | |
| 16023 
66561f6814bd
added string_of_thmref, selections, fact_index_of, valid_thms;
 wenzelm parents: 
15975diff
changeset | 317 | |
| 6091 | 318 | (* add_thms(s) *) | 
| 4853 | 319 | |
| 70428 | 320 | val add_thmss = | 
| 70430 | 321 | fold_map (fn ((b, thms), atts) => apply_facts official1 official2 (b, [(thms, atts)])); | 
| 5907 | 322 | |
| 70428 | 323 | fun add_thms args = | 
| 324 | add_thmss (map (apfst (apsnd single)) args) #>> map the_single; | |
| 12235 
5fa04fc9b254
Further restructuring of theorem naming functions.
 berghofe parents: 
12138diff
changeset | 325 | |
| 27683 | 326 | val add_thm = yield_singleton add_thms; | 
| 5907 | 327 | |
| 328 | ||
| 57929 
c5063c033a5a
tuned signature -- proper Local_Theory.add_thms_dynamic;
 wenzelm parents: 
57887diff
changeset | 329 | (* dynamic theorems *) | 
| 26488 
b497e3187ec7
eliminated destructive/critical theorem database;
 wenzelm parents: 
26471diff
changeset | 330 | |
| 57929 
c5063c033a5a
tuned signature -- proper Local_Theory.add_thms_dynamic;
 wenzelm parents: 
57887diff
changeset | 331 | fun add_thms_dynamic' context arg thy = | 
| 70544 | 332 | let val (name, facts') = Facts.add_dynamic context arg (facts_of thy) | 
| 333 | in (name, map_facts (K facts') thy) end; | |
| 57929 
c5063c033a5a
tuned signature -- proper Local_Theory.add_thms_dynamic;
 wenzelm parents: 
57887diff
changeset | 334 | |
| 
c5063c033a5a
tuned signature -- proper Local_Theory.add_thms_dynamic;
 wenzelm parents: 
57887diff
changeset | 335 | fun add_thms_dynamic arg thy = | 
| 
c5063c033a5a
tuned signature -- proper Local_Theory.add_thms_dynamic;
 wenzelm parents: 
57887diff
changeset | 336 | add_thms_dynamic' (Context.Theory thy) arg thy |> snd; | 
| 26488 
b497e3187ec7
eliminated destructive/critical theorem database;
 wenzelm parents: 
26471diff
changeset | 337 | |
| 
b497e3187ec7
eliminated destructive/critical theorem database;
 wenzelm parents: 
26471diff
changeset | 338 | |
| 27728 | 339 | (* note_thmss *) | 
| 5907 | 340 | |
| 67713 | 341 | fun note_thms kind ((b, more_atts), facts) thy = | 
| 12711 | 342 | let | 
| 28965 | 343 | val name = Sign.full_name thy b; | 
| 70424 | 344 | val facts' = facts |> map (apsnd (fn atts => surround (Thm.kind kind) (atts @ more_atts))); | 
| 70430 | 345 | val (thms', thy') = thy |> apply_facts official1 official2 (b, facts'); | 
| 70424 | 346 | in ((name, thms'), thy') end; | 
| 67713 | 347 | |
| 348 | val note_thmss = fold_map o note_thms; | |
| 12711 | 349 | |
| 5280 | 350 | |
| 62170 | 351 | (* old-style defs *) | 
| 4022 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 wenzelm parents: 
4013diff
changeset | 352 | |
| 4853 | 353 | 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 | 354 | |
| 62169 | 355 | fun add unchecked overloaded = fold_map (fn ((b, prop), atts) => fn 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 | 356 | let | 
| 62170 | 357 | val context = Defs.global_context thy; | 
| 61262 
7bd1eb4b056e
tuned signature: eliminated pointless type Context.pretty;
 wenzelm parents: 
61261diff
changeset | 358 | val ((_, def), thy') = Thm.add_def context 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 | 359 | 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 | 360 | |> 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 | 361 | |> 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 | 362 | |> Thm.varifyT_global; | 
| 70430 | 363 | in thy' |> apply_facts unnamed official2 (b, [([thm], atts)]) |>> the_single end); | 
| 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 | 364 | |
| 4853 | 365 | 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 | 366 | |
| 62169 | 367 | val add_defs = add false; | 
| 368 | val add_defs_unchecked = add true; | |
| 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 | 369 | |
| 4853 | 370 | end; | 
| 4022 
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
 wenzelm parents: 
4013diff
changeset | 371 | |
| 3987 | 372 | end; |