| author | ballarin |
| Mon, 11 Apr 2005 12:34:34 +0200 | |
| changeset 15696 | 1da4ce092c0b |
| parent 15624 | 484178635bd8 |
| child 15703 | 727ef1b8b3ee |
| permissions | -rw-r--r-- |
| 3987 | 1 |
(* Title: Pure/pure_thy.ML |
2 |
ID: $Id$ |
|
3 |
Author: Markus Wenzel, TU Muenchen |
|
4 |
||
| 5091 | 5 |
Theorem database, derived theory operations, and the ProtoPure theory. |
| 3987 | 6 |
*) |
7 |
||
|
4022
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
wenzelm
parents:
4013
diff
changeset
|
8 |
signature BASIC_PURE_THY = |
|
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
wenzelm
parents:
4013
diff
changeset
|
9 |
sig |
|
15456
956d6acacf89
Specific theorems in a named list of theorems can now be referred to
berghofe
parents:
15387
diff
changeset
|
10 |
type thmref |
| 5000 | 11 |
val print_theorems: theory -> unit |
12 |
val print_theory: theory -> unit |
|
|
15456
956d6acacf89
Specific theorems in a named list of theorems can now be referred to
berghofe
parents:
15387
diff
changeset
|
13 |
val get_thm: theory -> thmref -> thm |
|
956d6acacf89
Specific theorems in a named list of theorems can now be referred to
berghofe
parents:
15387
diff
changeset
|
14 |
val get_thms: theory -> thmref -> thm list |
|
956d6acacf89
Specific theorems in a named list of theorems can now be referred to
berghofe
parents:
15387
diff
changeset
|
15 |
val get_thmss: theory -> thmref list -> thm list |
|
4022
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
wenzelm
parents:
4013
diff
changeset
|
16 |
val thms_of: theory -> (string * thm) list |
| 5091 | 17 |
structure ProtoPure: |
18 |
sig |
|
19 |
val thy: theory |
|
20 |
val Goal_def: thm |
|
21 |
end |
|
| 4853 | 22 |
end; |
|
4022
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
wenzelm
parents:
4013
diff
changeset
|
23 |
|
| 3987 | 24 |
signature PURE_THY = |
25 |
sig |
|
|
4022
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
wenzelm
parents:
4013
diff
changeset
|
26 |
include BASIC_PURE_THY |
|
15456
956d6acacf89
Specific theorems in a named list of theorems can now be referred to
berghofe
parents:
15387
diff
changeset
|
27 |
val get_thm_closure: theory -> thmref -> thm |
|
956d6acacf89
Specific theorems in a named list of theorems can now be referred to
berghofe
parents:
15387
diff
changeset
|
28 |
val get_thms_closure: theory -> thmref -> thm list |
| 9564 | 29 |
val single_thm: string -> thm list -> thm |
|
15456
956d6acacf89
Specific theorems in a named list of theorems can now be referred to
berghofe
parents:
15387
diff
changeset
|
30 |
val select_thm: thmref -> thm list -> thm list |
| 6367 | 31 |
val cond_extern_thm_sg: Sign.sg -> string -> xstring |
| 13274 | 32 |
val thms_containing: theory -> string list * string list -> (string * thm list) list |
| 13646 | 33 |
val thms_containing_consts: theory -> string list -> (string * thm) list |
|
15387
24aff9e3de3f
fixed bug in find functions that I introduced some time ago.
nipkow
parents:
14981
diff
changeset
|
34 |
val find_matching_thms: (thm -> thm list) * (term -> term) |
|
24aff9e3de3f
fixed bug in find functions that I introduced some time ago.
nipkow
parents:
14981
diff
changeset
|
35 |
-> theory -> term -> (string * thm) list |
| 13646 | 36 |
val find_intros: theory -> term -> (string * thm) list |
|
13800
16136d2da0db
Moved find_intros_goal from goals.ML to pure_thy.ML
berghofe
parents:
13713
diff
changeset
|
37 |
val find_intros_goal : theory -> thm -> int -> (string * thm) list |
| 13646 | 38 |
val find_elims : theory -> term -> (string * thm) list |
| 12695 | 39 |
val hide_thms: bool -> string list -> theory -> theory |
| 6091 | 40 |
val store_thm: (bstring * thm) * theory attribute list -> theory -> theory * thm |
| 7405 | 41 |
val smart_store_thms: (bstring * thm list) -> thm list |
|
12138
7cad58fbc866
renamed open_smart_store_thms to smart_store_thms_open;
wenzelm
parents:
12123
diff
changeset
|
42 |
val smart_store_thms_open: (bstring * thm list) -> thm list |
| 7899 | 43 |
val forall_elim_var: int -> thm -> thm |
44 |
val forall_elim_vars: int -> thm -> thm |
|
|
8419
4770b1a12a93
add_thms, add_axioms, add_defs: return theorems as well;
wenzelm
parents:
8039
diff
changeset
|
45 |
val add_thms: ((bstring * thm) * theory attribute list) list -> theory -> theory * thm list |
| 12711 | 46 |
val add_thmss: ((bstring * thm list) * theory attribute list) list -> theory |
47 |
-> theory * thm list list |
|
|
15624
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15570
diff
changeset
|
48 |
val note_thmss: |
|
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15570
diff
changeset
|
49 |
theory attribute -> ((bstring * theory attribute list) * |
|
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15570
diff
changeset
|
50 |
(thmref * theory attribute list) list) list -> theory -> |
|
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15570
diff
changeset
|
51 |
theory * (bstring * thm list) list |
|
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15570
diff
changeset
|
52 |
val note_thmss_i: |
|
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15570
diff
changeset
|
53 |
theory attribute -> ((bstring * theory attribute list) * |
|
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15570
diff
changeset
|
54 |
(thm list * theory attribute list) list) list -> theory -> |
|
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15570
diff
changeset
|
55 |
theory * (bstring * thm list) list |
| 15696 | 56 |
val note_thmss_accesses: |
|
15624
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15570
diff
changeset
|
57 |
(string -> string list) -> |
|
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15570
diff
changeset
|
58 |
theory attribute -> ((bstring * theory attribute list) * |
|
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15570
diff
changeset
|
59 |
(thmref * theory attribute list) list) list -> theory -> |
|
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15570
diff
changeset
|
60 |
theory * (bstring * thm list) list |
| 15696 | 61 |
val note_thmss_accesses_i: |
|
15624
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15570
diff
changeset
|
62 |
(string -> string list) -> |
|
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15570
diff
changeset
|
63 |
theory attribute -> ((bstring * theory attribute list) * |
|
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15570
diff
changeset
|
64 |
(thm list * theory attribute list) list) list -> theory -> |
|
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15570
diff
changeset
|
65 |
theory * (bstring * thm list) list |
|
8419
4770b1a12a93
add_thms, add_axioms, add_defs: return theorems as well;
wenzelm
parents:
8039
diff
changeset
|
66 |
val add_axioms: ((bstring * string) * theory attribute list) list -> theory -> theory * thm list |
|
4770b1a12a93
add_thms, add_axioms, add_defs: return theorems as well;
wenzelm
parents:
8039
diff
changeset
|
67 |
val add_axioms_i: ((bstring * term) * theory attribute list) list -> theory -> theory * thm list |
| 12711 | 68 |
val add_axiomss: ((bstring * string list) * theory attribute list) list -> theory |
69 |
-> theory * thm list list |
|
70 |
val add_axiomss_i: ((bstring * term list) * theory attribute list) list -> theory |
|
71 |
-> theory * thm list list |
|
| 9318 | 72 |
val add_defs: bool -> ((bstring * string) * theory attribute list) list |
73 |
-> theory -> theory * thm list |
|
74 |
val add_defs_i: bool -> ((bstring * term) * theory attribute list) list |
|
75 |
-> theory -> theory * thm list |
|
76 |
val add_defss: bool -> ((bstring * string list) * theory attribute list) list |
|
77 |
-> theory -> theory * thm list list |
|
78 |
val add_defss_i: bool -> ((bstring * term list) * theory attribute list) list |
|
79 |
-> theory -> theory * thm list list |
|
|
4963
38aa2d56e28c
added get_name, put_name, global_path, local_path, begin_theory,
wenzelm
parents:
4933
diff
changeset
|
80 |
val get_name: theory -> string |
|
38aa2d56e28c
added get_name, put_name, global_path, local_path, begin_theory,
wenzelm
parents:
4933
diff
changeset
|
81 |
val put_name: string -> theory -> theory |
|
38aa2d56e28c
added get_name, put_name, global_path, local_path, begin_theory,
wenzelm
parents:
4933
diff
changeset
|
82 |
val global_path: theory -> theory |
|
38aa2d56e28c
added get_name, put_name, global_path, local_path, begin_theory,
wenzelm
parents:
4933
diff
changeset
|
83 |
val local_path: theory -> theory |
|
38aa2d56e28c
added get_name, put_name, global_path, local_path, begin_theory,
wenzelm
parents:
4933
diff
changeset
|
84 |
val begin_theory: string -> theory list -> theory |
|
38aa2d56e28c
added get_name, put_name, global_path, local_path, begin_theory,
wenzelm
parents:
4933
diff
changeset
|
85 |
val end_theory: theory -> theory |
| 6682 | 86 |
val checkpoint: theory -> theory |
|
4922
03b81b6e1baa
added thms_closure: theory -> xstring -> tthm list option;
wenzelm
parents:
4853
diff
changeset
|
87 |
val add_typedecls: (bstring * string list * mixfix) list -> theory -> theory |
| 3987 | 88 |
end; |
89 |
||
90 |
structure PureThy: PURE_THY = |
|
91 |
struct |
|
92 |
||
93 |
||
|
4922
03b81b6e1baa
added thms_closure: theory -> xstring -> tthm list option;
wenzelm
parents:
4853
diff
changeset
|
94 |
(*** theorem database ***) |
| 3987 | 95 |
|
|
4922
03b81b6e1baa
added thms_closure: theory -> xstring -> tthm list option;
wenzelm
parents:
4853
diff
changeset
|
96 |
(** data kind 'Pure/theorems' **) |
| 3987 | 97 |
|
| 5005 | 98 |
structure TheoremsDataArgs = |
99 |
struct |
|
100 |
val name = "Pure/theorems"; |
|
| 3987 | 101 |
|
| 5005 | 102 |
type T = |
103 |
{space: NameSpace.T,
|
|
| 6091 | 104 |
thms_tab: thm list Symtab.table, |
| 13274 | 105 |
index: FactIndex.T} ref; |
| 3987 | 106 |
|
| 4853 | 107 |
fun mk_empty _ = |
| 13274 | 108 |
ref {space = NameSpace.empty, thms_tab = Symtab.empty, index = FactIndex.empty}: T;
|
| 3987 | 109 |
|
| 5005 | 110 |
val empty = mk_empty (); |
| 6547 | 111 |
fun copy (ref x) = ref x; |
| 5005 | 112 |
val prep_ext = mk_empty; |
113 |
val merge = mk_empty; |
|
114 |
||
| 13274 | 115 |
fun pretty sg (ref {space, thms_tab, index = _}) =
|
| 4853 | 116 |
let |
| 10008 | 117 |
val prt_thm = Display.pretty_thm_sg sg; |
| 4853 | 118 |
fun prt_thms (name, [th]) = |
119 |
Pretty.block [Pretty.str (name ^ ":"), Pretty.brk 1, prt_thm th] |
|
120 |
| prt_thms (name, ths) = Pretty.big_list (name ^ ":") (map prt_thm ths); |
|
| 3987 | 121 |
|
| 6846 | 122 |
val thmss = NameSpace.cond_extern_table space thms_tab; |
| 9215 | 123 |
in Pretty.big_list "theorems:" (map prt_thms thmss) end; |
| 8720 | 124 |
|
| 9215 | 125 |
fun print sg data = Pretty.writeln (pretty sg data); |
| 3987 | 126 |
end; |
127 |
||
| 5005 | 128 |
structure TheoremsData = TheoryDataFun(TheoremsDataArgs); |
129 |
val get_theorems_sg = TheoremsData.get_sg; |
|
130 |
val get_theorems = TheoremsData.get; |
|
131 |
||
| 6367 | 132 |
val cond_extern_thm_sg = NameSpace.cond_extern o #space o ! o get_theorems_sg; |
133 |
||
| 3987 | 134 |
|
| 5000 | 135 |
(* print theory *) |
| 3987 | 136 |
|
| 5005 | 137 |
val print_theorems = TheoremsData.print; |
| 8720 | 138 |
|
| 5000 | 139 |
fun print_theory thy = |
| 9215 | 140 |
Display.pretty_full_theory thy @ |
141 |
[TheoremsDataArgs.pretty (Theory.sign_of thy) (get_theorems thy)] |
|
| 8720 | 142 |
|> Pretty.chunks |> Pretty.writeln; |
|
4022
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
wenzelm
parents:
4013
diff
changeset
|
143 |
|
| 3987 | 144 |
|
145 |
||
|
4022
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
wenzelm
parents:
4013
diff
changeset
|
146 |
(** retrieve theorems **) |
| 3987 | 147 |
|
|
15456
956d6acacf89
Specific theorems in a named list of theorems can now be referred to
berghofe
parents:
15387
diff
changeset
|
148 |
type thmref = xstring * int list option; |
|
956d6acacf89
Specific theorems in a named list of theorems can now be referred to
berghofe
parents:
15387
diff
changeset
|
149 |
|
| 9564 | 150 |
(* selections *) |
151 |
||
| 15531 | 152 |
fun the_thms _ (SOME thms) = thms |
153 |
| the_thms name NONE = error ("Unknown theorem(s) " ^ quote name);
|
|
| 4037 | 154 |
|
| 9564 | 155 |
fun single_thm _ [thm] = thm |
156 |
| single_thm name _ = error ("Single theorem expected " ^ quote name);
|
|
157 |
||
| 15531 | 158 |
fun select_thm (s, NONE) xs = xs |
159 |
| select_thm (s, SOME is) xs = map |
|
| 15570 | 160 |
(fn i => (if i < 1 then raise Subscript else List.nth (xs, i-1)) handle Subscript => |
|
15456
956d6acacf89
Specific theorems in a named list of theorems can now be referred to
berghofe
parents:
15387
diff
changeset
|
161 |
error ("Bad subscript " ^ string_of_int i ^ " for " ^ quote s)) is;
|
|
956d6acacf89
Specific theorems in a named list of theorems can now be referred to
berghofe
parents:
15387
diff
changeset
|
162 |
|
| 9564 | 163 |
|
| 9808 | 164 |
(* get_thm(s)_closure -- statically scoped versions *) |
| 9564 | 165 |
|
166 |
(*beware of proper order of evaluation!*) |
|
|
4922
03b81b6e1baa
added thms_closure: theory -> xstring -> tthm list option;
wenzelm
parents:
4853
diff
changeset
|
167 |
|
| 9564 | 168 |
fun lookup_thms thy = |
169 |
let |
|
170 |
val sg_ref = Sign.self_ref (Theory.sign_of thy); |
|
171 |
val ref {space, thms_tab, ...} = get_theorems thy;
|
|
172 |
in |
|
173 |
fn name => |
|
| 15570 | 174 |
Option.map (map (Thm.transfer_sg (Sign.deref sg_ref))) (*semi-dynamic identity*) |
| 9564 | 175 |
(Symtab.lookup (thms_tab, NameSpace.intern space name)) (*static content*) |
176 |
end; |
|
| 3987 | 177 |
|
| 9564 | 178 |
fun get_thms_closure thy = |
179 |
let val closures = map lookup_thms (thy :: Theory.ancestors_of thy) |
|
|
15456
956d6acacf89
Specific theorems in a named list of theorems can now be referred to
berghofe
parents:
15387
diff
changeset
|
180 |
in fn namei as (name, _) => select_thm namei |
|
956d6acacf89
Specific theorems in a named list of theorems can now be referred to
berghofe
parents:
15387
diff
changeset
|
181 |
(the_thms name (get_first (fn f => f name) closures)) |
|
956d6acacf89
Specific theorems in a named list of theorems can now be referred to
berghofe
parents:
15387
diff
changeset
|
182 |
end; |
| 9564 | 183 |
|
| 9808 | 184 |
fun get_thm_closure thy = |
185 |
let val get = get_thms_closure thy |
|
|
15456
956d6acacf89
Specific theorems in a named list of theorems can now be referred to
berghofe
parents:
15387
diff
changeset
|
186 |
in fn namei as (name, _) => single_thm name (get namei) end; |
| 9808 | 187 |
|
| 9564 | 188 |
|
189 |
(* get_thm etc. *) |
|
190 |
||
|
15456
956d6acacf89
Specific theorems in a named list of theorems can now be referred to
berghofe
parents:
15387
diff
changeset
|
191 |
fun get_thms theory (namei as (name, _)) = |
| 9564 | 192 |
get_first (fn thy => lookup_thms thy name) (theory :: Theory.ancestors_of theory) |
|
15456
956d6acacf89
Specific theorems in a named list of theorems can now be referred to
berghofe
parents:
15387
diff
changeset
|
193 |
|> the_thms name |> select_thm namei |> map (Thm.transfer theory); |
|
4022
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
wenzelm
parents:
4013
diff
changeset
|
194 |
|
| 15570 | 195 |
fun get_thmss thy names = List.concat (map (get_thms thy) names); |
|
15456
956d6acacf89
Specific theorems in a named list of theorems can now be referred to
berghofe
parents:
15387
diff
changeset
|
196 |
fun get_thm thy (namei as (name, _)) = single_thm name (get_thms thy namei); |
| 4783 | 197 |
|
|
4022
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
wenzelm
parents:
4013
diff
changeset
|
198 |
|
|
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
wenzelm
parents:
4013
diff
changeset
|
199 |
(* thms_of *) |
|
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
wenzelm
parents:
4013
diff
changeset
|
200 |
|
|
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
wenzelm
parents:
4013
diff
changeset
|
201 |
fun thms_of thy = |
|
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
wenzelm
parents:
4013
diff
changeset
|
202 |
let val ref {thms_tab, ...} = get_theorems thy in
|
| 15570 | 203 |
map (fn th => (Thm.name_of_thm th, th)) (List.concat (map snd (Symtab.dest thms_tab))) |
|
4022
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
wenzelm
parents:
4013
diff
changeset
|
204 |
end; |
|
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
wenzelm
parents:
4013
diff
changeset
|
205 |
|
| 3987 | 206 |
|
| 13274 | 207 |
(* thms_containing *) |
| 3987 | 208 |
|
| 13274 | 209 |
fun thms_containing thy idx = |
210 |
let |
|
211 |
fun valid (name, ths) = |
|
| 15531 | 212 |
(case try (transform_error (get_thms thy)) (name, NONE) of |
213 |
NONE => false |
|
214 |
| SOME ths' => Library.equal_lists Thm.eq_thm (ths, ths')); |
|
| 13274 | 215 |
in |
216 |
(thy :: Theory.ancestors_of thy) |
|
| 15570 | 217 |
|> map (gen_distinct eq_fst o List.filter valid o FactIndex.find idx o #index o ! o get_theorems) |
218 |
|> List.concat |
|
| 13274 | 219 |
end; |
|
4022
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
wenzelm
parents:
4013
diff
changeset
|
220 |
|
| 13646 | 221 |
fun thms_containing_consts thy consts = |
| 15570 | 222 |
thms_containing thy (consts, []) |> map #2 |> List.concat |
| 13646 | 223 |
|> map (fn th => (Thm.name_of_thm th, th)) |
224 |
||
225 |
(* intro/elim theorems *) |
|
226 |
||
227 |
(* intro: given a goal state, find a suitable intro rule for some subgoal *) |
|
228 |
(* elim: given a theorem thm, |
|
229 |
find a theorem whose major premise eliminates the conclusion of thm *) |
|
230 |
||
| 15531 | 231 |
fun top_const t = (case head_of t of Const (c, _) => SOME c | _ => NONE); |
| 13646 | 232 |
|
|
15387
24aff9e3de3f
fixed bug in find functions that I introduced some time ago.
nipkow
parents:
14981
diff
changeset
|
233 |
(* This is a hack to remove the Trueprop constant that most logics use *) |
|
24aff9e3de3f
fixed bug in find functions that I introduced some time ago.
nipkow
parents:
14981
diff
changeset
|
234 |
fun rem_top (_ $ t) = t |
|
24aff9e3de3f
fixed bug in find functions that I introduced some time ago.
nipkow
parents:
14981
diff
changeset
|
235 |
| rem_top _ = Bound 0 (* does not match anything *) |
| 13646 | 236 |
|
237 |
(*returns all those named_thms whose subterm extracted by extract can be |
|
238 |
instantiated to obj; the list is sorted according to the number of premises |
|
239 |
and the size of the required substitution.*) |
|
|
15387
24aff9e3de3f
fixed bug in find functions that I introduced some time ago.
nipkow
parents:
14981
diff
changeset
|
240 |
fun select_match(c,obj, signobj, named_thms, (extract_thms,extract_term)) = |
| 13646 | 241 |
let val tsig = Sign.tsig_of signobj |
242 |
fun matches prop = |
|
|
15387
24aff9e3de3f
fixed bug in find functions that I introduced some time ago.
nipkow
parents:
14981
diff
changeset
|
243 |
let val pat = extract_term prop |
|
24aff9e3de3f
fixed bug in find functions that I introduced some time ago.
nipkow
parents:
14981
diff
changeset
|
244 |
in case head_of pat of |
|
24aff9e3de3f
fixed bug in find functions that I introduced some time ago.
nipkow
parents:
14981
diff
changeset
|
245 |
Const(d,_) => c=d andalso Pattern.matches tsig (pat,obj) |
|
24aff9e3de3f
fixed bug in find functions that I introduced some time ago.
nipkow
parents:
14981
diff
changeset
|
246 |
| _ => false |
|
24aff9e3de3f
fixed bug in find functions that I introduced some time ago.
nipkow
parents:
14981
diff
changeset
|
247 |
end |
| 13646 | 248 |
|
249 |
fun substsize prop = |
|
|
15387
24aff9e3de3f
fixed bug in find functions that I introduced some time ago.
nipkow
parents:
14981
diff
changeset
|
250 |
let val pat = extract_term prop |
| 13646 | 251 |
val (_,subst) = Pattern.match tsig (pat,obj) |
| 15570 | 252 |
in Library.foldl op+ (0, map (size_of_term o snd) subst) end |
| 13646 | 253 |
|
254 |
fun thm_ord ((p0,s0,_),(p1,s1,_)) = |
|
255 |
prod_ord (int_ord o pairself (fn 0 => 0 | x => 1)) int_ord ((p0,s0),(p1,s1)); |
|
256 |
||
257 |
fun select((p as (_,thm))::named_thms, sels) = |
|
|
15387
24aff9e3de3f
fixed bug in find functions that I introduced some time ago.
nipkow
parents:
14981
diff
changeset
|
258 |
let |
|
24aff9e3de3f
fixed bug in find functions that I introduced some time ago.
nipkow
parents:
14981
diff
changeset
|
259 |
fun sel(thm::thms,sels) = |
|
24aff9e3de3f
fixed bug in find functions that I introduced some time ago.
nipkow
parents:
14981
diff
changeset
|
260 |
let val {prop, ...} = rep_thm thm
|
|
24aff9e3de3f
fixed bug in find functions that I introduced some time ago.
nipkow
parents:
14981
diff
changeset
|
261 |
in if matches prop |
|
24aff9e3de3f
fixed bug in find functions that I introduced some time ago.
nipkow
parents:
14981
diff
changeset
|
262 |
then (nprems_of thm,substsize prop,p)::sels |
|
24aff9e3de3f
fixed bug in find functions that I introduced some time ago.
nipkow
parents:
14981
diff
changeset
|
263 |
else sel(thms,sels) |
|
24aff9e3de3f
fixed bug in find functions that I introduced some time ago.
nipkow
parents:
14981
diff
changeset
|
264 |
end |
|
24aff9e3de3f
fixed bug in find functions that I introduced some time ago.
nipkow
parents:
14981
diff
changeset
|
265 |
| sel([],sels) = sels |
|
24aff9e3de3f
fixed bug in find functions that I introduced some time ago.
nipkow
parents:
14981
diff
changeset
|
266 |
val {sign, ...} = rep_thm thm
|
|
24aff9e3de3f
fixed bug in find functions that I introduced some time ago.
nipkow
parents:
14981
diff
changeset
|
267 |
in select(named_thms,if Sign.subsig(sign, signobj) |
|
24aff9e3de3f
fixed bug in find functions that I introduced some time ago.
nipkow
parents:
14981
diff
changeset
|
268 |
then sel(extract_thms thm,sels) |
|
24aff9e3de3f
fixed bug in find functions that I introduced some time ago.
nipkow
parents:
14981
diff
changeset
|
269 |
else sels) |
|
24aff9e3de3f
fixed bug in find functions that I introduced some time ago.
nipkow
parents:
14981
diff
changeset
|
270 |
end |
| 13646 | 271 |
| select([],sels) = sels |
272 |
||
273 |
in map (fn (_,_,t) => t) (sort thm_ord (select(named_thms, []))) end; |
|
274 |
||
|
15387
24aff9e3de3f
fixed bug in find functions that I introduced some time ago.
nipkow
parents:
14981
diff
changeset
|
275 |
fun find_matching_thms extract thy prop = |
| 15531 | 276 |
(case top_const prop of NONE => [] |
277 |
| SOME c => let val thms = thms_containing_consts thy [c] |
|
|
15387
24aff9e3de3f
fixed bug in find functions that I introduced some time ago.
nipkow
parents:
14981
diff
changeset
|
278 |
in select_match(c,prop,Theory.sign_of thy,thms,extract) end) |
| 13646 | 279 |
|
|
15387
24aff9e3de3f
fixed bug in find functions that I introduced some time ago.
nipkow
parents:
14981
diff
changeset
|
280 |
val find_intros = |
|
24aff9e3de3f
fixed bug in find functions that I introduced some time ago.
nipkow
parents:
14981
diff
changeset
|
281 |
find_matching_thms (single, rem_top o Logic.strip_imp_concl) |
| 13646 | 282 |
|
|
13800
16136d2da0db
Moved find_intros_goal from goals.ML to pure_thy.ML
berghofe
parents:
13713
diff
changeset
|
283 |
fun find_intros_goal thy st i = |
|
15387
24aff9e3de3f
fixed bug in find functions that I introduced some time ago.
nipkow
parents:
14981
diff
changeset
|
284 |
find_intros thy (rem_top(Logic.concl_of_goal (prop_of st) i)); |
|
13800
16136d2da0db
Moved find_intros_goal from goals.ML to pure_thy.ML
berghofe
parents:
13713
diff
changeset
|
285 |
|
|
15387
24aff9e3de3f
fixed bug in find functions that I introduced some time ago.
nipkow
parents:
14981
diff
changeset
|
286 |
val find_elims = find_matching_thms |
|
24aff9e3de3f
fixed bug in find functions that I introduced some time ago.
nipkow
parents:
14981
diff
changeset
|
287 |
(fn thm => if Thm.no_prems thm then [] else [thm], |
|
24aff9e3de3f
fixed bug in find functions that I introduced some time ago.
nipkow
parents:
14981
diff
changeset
|
288 |
rem_top o hd o Logic.strip_imp_prems) |
|
4022
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
wenzelm
parents:
4013
diff
changeset
|
289 |
|
|
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
wenzelm
parents:
4013
diff
changeset
|
290 |
|
|
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
wenzelm
parents:
4013
diff
changeset
|
291 |
(** store theorems **) (*DESTRUCTIVE*) |
| 3987 | 292 |
|
| 12695 | 293 |
(* hiding -- affects current theory node only! *) |
294 |
||
| 13424 | 295 |
fun hide_thms fully names thy = |
| 12695 | 296 |
let |
| 13274 | 297 |
val r as ref {space, thms_tab, index} = get_theorems thy;
|
| 13424 | 298 |
val space' = NameSpace.hide fully (space, names); |
| 13274 | 299 |
in r := {space = space', thms_tab = thms_tab, index = index}; thy end;
|
| 12695 | 300 |
|
301 |
||
| 4853 | 302 |
(* naming *) |
303 |
||
| 11998 | 304 |
fun gen_names j len name = |
305 |
map (fn i => name ^ "_" ^ string_of_int i) (j+1 upto j+len); |
|
| 4853 | 306 |
|
| 11998 | 307 |
fun name_multi name xs = gen_names 0 (length xs) name ~~ xs; |
|
12235
5fa04fc9b254
Further restructuring of theorem naming functions.
berghofe
parents:
12138
diff
changeset
|
308 |
|
|
12872
0855c3ab2047
Theorems are only "pre-named" if the do not already have names.
berghofe
parents:
12711
diff
changeset
|
309 |
fun name_thm pre (p as (_, thm)) = |
|
0855c3ab2047
Theorems are only "pre-named" if the do not already have names.
berghofe
parents:
12711
diff
changeset
|
310 |
if Thm.name_of_thm thm <> "" andalso pre then thm else Thm.name_thm p; |
|
0855c3ab2047
Theorems are only "pre-named" if the do not already have names.
berghofe
parents:
12711
diff
changeset
|
311 |
|
|
0855c3ab2047
Theorems are only "pre-named" if the do not already have names.
berghofe
parents:
12711
diff
changeset
|
312 |
fun name_thms pre name [x] = [name_thm pre (name, x)] |
|
0855c3ab2047
Theorems are only "pre-named" if the do not already have names.
berghofe
parents:
12711
diff
changeset
|
313 |
| name_thms pre name xs = map (name_thm pre) (name_multi name xs); |
|
12235
5fa04fc9b254
Further restructuring of theorem naming functions.
berghofe
parents:
12138
diff
changeset
|
314 |
|
|
5fa04fc9b254
Further restructuring of theorem naming functions.
berghofe
parents:
12138
diff
changeset
|
315 |
fun name_thmss name xs = (case filter_out (null o fst) xs of |
|
12872
0855c3ab2047
Theorems are only "pre-named" if the do not already have names.
berghofe
parents:
12711
diff
changeset
|
316 |
[([x], z)] => [([name_thm true (name, x)], z)] |
|
12235
5fa04fc9b254
Further restructuring of theorem naming functions.
berghofe
parents:
12138
diff
changeset
|
317 |
| _ => snd (foldl_map (fn (i, (ys, z)) => (i + length ys, |
|
12872
0855c3ab2047
Theorems are only "pre-named" if the do not already have names.
berghofe
parents:
12711
diff
changeset
|
318 |
(map (name_thm true) (gen_names i (length ys) name ~~ ys), z))) (0, xs))); |
| 4853 | 319 |
|
320 |
||
| 11998 | 321 |
(* enter_thms *) |
| 4853 | 322 |
|
|
7470
9f67ca1e03dc
eliminated default_name (thms no longer stored for name "");
wenzelm
parents:
7405
diff
changeset
|
323 |
fun warn_overwrite name = warning ("Replaced old copy of theorems " ^ quote name);
|
|
9f67ca1e03dc
eliminated default_name (thms no longer stored for name "");
wenzelm
parents:
7405
diff
changeset
|
324 |
fun warn_same name = warning ("Theorem database already contains a copy of " ^ quote name);
|
| 3987 | 325 |
|
|
15624
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15570
diff
changeset
|
326 |
fun gen_enter_thms _ _ _ _ _ app_att thy ("", thms) = app_att (thy, thms)
|
|
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15570
diff
changeset
|
327 |
| gen_enter_thms full acc sg pre_name post_name app_att thy (bname, thms) = |
|
7470
9f67ca1e03dc
eliminated default_name (thms no longer stored for name "");
wenzelm
parents:
7405
diff
changeset
|
328 |
let |
|
15624
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15570
diff
changeset
|
329 |
val name = full sg bname; |
| 11998 | 330 |
val (thy', thms') = app_att (thy, pre_name name thms); |
331 |
val named_thms = post_name name thms'; |
|
| 3987 | 332 |
|
| 13274 | 333 |
val r as ref {space, thms_tab, index} = get_theorems_sg sg;
|
|
15624
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15570
diff
changeset
|
334 |
val space' = NameSpace.extend' acc (space, [name]); |
|
7470
9f67ca1e03dc
eliminated default_name (thms no longer stored for name "");
wenzelm
parents:
7405
diff
changeset
|
335 |
val thms_tab' = Symtab.update ((name, named_thms), thms_tab); |
| 13274 | 336 |
val index' = FactIndex.add (K false) (index, (name, named_thms)); |
337 |
in |
|
338 |
(case Symtab.lookup (thms_tab, name) of |
|
| 15531 | 339 |
NONE => () |
340 |
| SOME thms' => |
|
| 13274 | 341 |
if Library.equal_lists Thm.eq_thm (thms', named_thms) then warn_same name |
342 |
else warn_overwrite name); |
|
343 |
r := {space = space', thms_tab = thms_tab', index = index'};
|
|
| 11998 | 344 |
(thy', named_thms) |
345 |
end; |
|
| 3987 | 346 |
|
|
15624
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15570
diff
changeset
|
347 |
fun enter_thms sg = gen_enter_thms Sign.full_name NameSpace.accesses sg; |
| 4853 | 348 |
|
| 6091 | 349 |
(* add_thms(s) *) |
| 4853 | 350 |
|
|
12235
5fa04fc9b254
Further restructuring of theorem naming functions.
berghofe
parents:
12138
diff
changeset
|
351 |
fun add_thms_atts pre_name ((bname, thms), atts) thy = |
|
12872
0855c3ab2047
Theorems are only "pre-named" if the do not already have names.
berghofe
parents:
12711
diff
changeset
|
352 |
enter_thms (Theory.sign_of thy) pre_name (name_thms false) |
| 11998 | 353 |
(Thm.applys_attributes o rpair atts) thy (bname, thms); |
| 4853 | 354 |
|
|
12235
5fa04fc9b254
Further restructuring of theorem naming functions.
berghofe
parents:
12138
diff
changeset
|
355 |
fun gen_add_thmss pre_name args theory = |
|
5fa04fc9b254
Further restructuring of theorem naming functions.
berghofe
parents:
12138
diff
changeset
|
356 |
foldl_map (fn (thy, arg) => add_thms_atts pre_name arg thy) (theory, args); |
| 5907 | 357 |
|
|
12235
5fa04fc9b254
Further restructuring of theorem naming functions.
berghofe
parents:
12138
diff
changeset
|
358 |
fun gen_add_thms pre_name args = |
|
5fa04fc9b254
Further restructuring of theorem naming functions.
berghofe
parents:
12138
diff
changeset
|
359 |
apsnd (map hd) o gen_add_thmss pre_name (map (apfst (apsnd single)) args); |
|
5fa04fc9b254
Further restructuring of theorem naming functions.
berghofe
parents:
12138
diff
changeset
|
360 |
|
|
12872
0855c3ab2047
Theorems are only "pre-named" if the do not already have names.
berghofe
parents:
12711
diff
changeset
|
361 |
val add_thmss = gen_add_thmss (name_thms true); |
|
0855c3ab2047
Theorems are only "pre-named" if the do not already have names.
berghofe
parents:
12711
diff
changeset
|
362 |
val add_thms = gen_add_thms (name_thms true); |
| 5907 | 363 |
|
364 |
||
| 14564 | 365 |
(* note_thmss(_i) *) |
| 5907 | 366 |
|
| 9192 | 367 |
local |
| 12711 | 368 |
|
|
15624
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15570
diff
changeset
|
369 |
fun gen_note_thss enter get kind_att (thy, ((bname, more_atts), ths_atts)) = |
| 12711 | 370 |
let |
371 |
fun app (x, (ths, atts)) = Thm.applys_attributes ((x, ths), atts); |
|
|
15624
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15570
diff
changeset
|
372 |
val (thy', thms) = enter (Theory.sign_of thy) |
| 15570 | 373 |
name_thmss (name_thms false) (apsnd List.concat o foldl_map app) thy |
| 12711 | 374 |
(bname, map (fn (ths, atts) => (get thy ths, atts @ more_atts @ [kind_att])) ths_atts); |
375 |
in (thy', (bname, thms)) end; |
|
376 |
||
|
15624
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15570
diff
changeset
|
377 |
fun gen_note_thmss enter get kind_att args thy = |
|
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15570
diff
changeset
|
378 |
foldl_map (gen_note_thss enter get kind_att) (thy, args); |
| 12711 | 379 |
|
| 9192 | 380 |
in |
| 12711 | 381 |
|
|
15624
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15570
diff
changeset
|
382 |
(* if path is set, only permit unqualified names *) |
|
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15570
diff
changeset
|
383 |
|
|
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15570
diff
changeset
|
384 |
val note_thmss = gen_note_thmss enter_thms get_thms; |
|
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15570
diff
changeset
|
385 |
val note_thmss_i = gen_note_thmss enter_thms (K I); |
|
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15570
diff
changeset
|
386 |
|
|
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15570
diff
changeset
|
387 |
(* always permit qualified names, |
|
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15570
diff
changeset
|
388 |
clients may specify non-standard access policy *) |
|
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15570
diff
changeset
|
389 |
|
| 15696 | 390 |
fun note_thmss_accesses acc = |
|
15624
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15570
diff
changeset
|
391 |
gen_note_thmss (gen_enter_thms Sign.full_name' acc) get_thms; |
| 15696 | 392 |
fun note_thmss_accesses_i acc = |
|
15624
484178635bd8
Further work on interpretation commands. New command `interpret' for
ballarin
parents:
15570
diff
changeset
|
393 |
gen_note_thmss (gen_enter_thms Sign.full_name' acc) (K I); |
| 12711 | 394 |
|
| 9192 | 395 |
end; |
| 5280 | 396 |
|
397 |
||
| 6091 | 398 |
(* store_thm *) |
| 5280 | 399 |
|
| 11998 | 400 |
fun store_thm ((bname, thm), atts) thy = |
|
12872
0855c3ab2047
Theorems are only "pre-named" if the do not already have names.
berghofe
parents:
12711
diff
changeset
|
401 |
let val (thy', [th']) = add_thms_atts (name_thms true) ((bname, [thm]), atts) thy |
| 5280 | 402 |
in (thy', th') end; |
| 3987 | 403 |
|
404 |
||
| 7405 | 405 |
(* smart_store_thms *) |
| 3987 | 406 |
|
|
12235
5fa04fc9b254
Further restructuring of theorem naming functions.
berghofe
parents:
12138
diff
changeset
|
407 |
fun gen_smart_store_thms _ (name, []) = |
|
11516
a0633bdcd015
Added equality axioms and initialization of proof term package.
berghofe
parents:
10667
diff
changeset
|
408 |
error ("Cannot store empty list of theorems: " ^ quote name)
|
|
12235
5fa04fc9b254
Further restructuring of theorem naming functions.
berghofe
parents:
12138
diff
changeset
|
409 |
| gen_smart_store_thms name_thm (name, [thm]) = |
|
12872
0855c3ab2047
Theorems are only "pre-named" if the do not already have names.
berghofe
parents:
12711
diff
changeset
|
410 |
snd (enter_thms (Thm.sign_of_thm thm) (name_thm true) (name_thm false) |
|
0855c3ab2047
Theorems are only "pre-named" if the do not already have names.
berghofe
parents:
12711
diff
changeset
|
411 |
I () (name, [thm])) |
|
12235
5fa04fc9b254
Further restructuring of theorem naming functions.
berghofe
parents:
12138
diff
changeset
|
412 |
| gen_smart_store_thms name_thm (name, thms) = |
| 7405 | 413 |
let |
414 |
val merge_sg = Sign.merge_refs o apsnd (Sign.self_ref o Thm.sign_of_thm); |
|
| 15570 | 415 |
val sg_ref = Library.foldl merge_sg (Sign.self_ref (Thm.sign_of_thm (hd thms)), tl thms); |
|
12872
0855c3ab2047
Theorems are only "pre-named" if the do not already have names.
berghofe
parents:
12711
diff
changeset
|
416 |
in snd (enter_thms (Sign.deref sg_ref) (name_thm true) (name_thm false) |
|
0855c3ab2047
Theorems are only "pre-named" if the do not already have names.
berghofe
parents:
12711
diff
changeset
|
417 |
I () (name, thms)) |
|
0855c3ab2047
Theorems are only "pre-named" if the do not already have names.
berghofe
parents:
12711
diff
changeset
|
418 |
end; |
|
11516
a0633bdcd015
Added equality axioms and initialization of proof term package.
berghofe
parents:
10667
diff
changeset
|
419 |
|
|
12235
5fa04fc9b254
Further restructuring of theorem naming functions.
berghofe
parents:
12138
diff
changeset
|
420 |
val smart_store_thms = gen_smart_store_thms name_thms; |
|
12872
0855c3ab2047
Theorems are only "pre-named" if the do not already have names.
berghofe
parents:
12711
diff
changeset
|
421 |
val smart_store_thms_open = gen_smart_store_thms (K (K I)); |
| 3987 | 422 |
|
423 |
||
| 7899 | 424 |
(* forall_elim_vars (belongs to drule.ML) *) |
425 |
||
| 13713 | 426 |
(*Replace outermost quantified variable by Var of given index.*) |
| 7899 | 427 |
fun forall_elim_var i th = |
428 |
let val {prop,sign,...} = rep_thm th
|
|
429 |
in case prop of |
|
| 13713 | 430 |
Const ("all", _) $ Abs (a, T, _) =>
|
431 |
let val used = map (fst o fst) |
|
| 15570 | 432 |
(List.filter (equal i o snd o fst) (Term.add_vars ([], prop))) |
| 13713 | 433 |
in forall_elim (cterm_of sign (Var ((variant used a, i), T))) th end |
434 |
| _ => raise THM ("forall_elim_var", i, [th])
|
|
| 7899 | 435 |
end; |
436 |
||
437 |
(*Repeat forall_elim_var until all outer quantifiers are removed*) |
|
438 |
fun forall_elim_vars i th = |
|
439 |
forall_elim_vars i (forall_elim_var i th) |
|
440 |
handle THM _ => th; |
|
441 |
||
442 |
||
|
4022
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
wenzelm
parents:
4013
diff
changeset
|
443 |
(* store axioms as theorems *) |
|
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
wenzelm
parents:
4013
diff
changeset
|
444 |
|
| 4853 | 445 |
local |
| 7899 | 446 |
fun get_axs thy named_axs = |
447 |
map (forall_elim_vars 0 o Thm.get_axiom thy o fst) named_axs; |
|
| 7753 | 448 |
|
|
8419
4770b1a12a93
add_thms, add_axioms, add_defs: return theorems as well;
wenzelm
parents:
8039
diff
changeset
|
449 |
fun add_single add (thy, ((name, ax), atts)) = |
| 4853 | 450 |
let |
| 11998 | 451 |
val named_ax = [(name, ax)]; |
| 7753 | 452 |
val thy' = add named_ax thy; |
453 |
val thm = hd (get_axs thy' named_ax); |
|
|
12235
5fa04fc9b254
Further restructuring of theorem naming functions.
berghofe
parents:
12138
diff
changeset
|
454 |
in apsnd hd (gen_add_thms (K I) [((name, thm), atts)] thy') end; |
| 7753 | 455 |
|
|
8419
4770b1a12a93
add_thms, add_axioms, add_defs: return theorems as well;
wenzelm
parents:
8039
diff
changeset
|
456 |
fun add_multi add (thy, ((name, axs), atts)) = |
| 7753 | 457 |
let |
458 |
val named_axs = name_multi name axs; |
|
| 4853 | 459 |
val thy' = add named_axs thy; |
| 7753 | 460 |
val thms = get_axs thy' named_axs; |
|
12235
5fa04fc9b254
Further restructuring of theorem naming functions.
berghofe
parents:
12138
diff
changeset
|
461 |
in apsnd hd (gen_add_thmss (K I) [((name, thms), atts)] thy') end; |
|
4022
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
wenzelm
parents:
4013
diff
changeset
|
462 |
|
|
8419
4770b1a12a93
add_thms, add_axioms, add_defs: return theorems as well;
wenzelm
parents:
8039
diff
changeset
|
463 |
fun add_singles add args thy = foldl_map (add_single add) (thy, args); |
|
4770b1a12a93
add_thms, add_axioms, add_defs: return theorems as well;
wenzelm
parents:
8039
diff
changeset
|
464 |
fun add_multis add args thy = foldl_map (add_multi add) (thy, args); |
| 4853 | 465 |
in |
| 7753 | 466 |
val add_axioms = add_singles Theory.add_axioms; |
467 |
val add_axioms_i = add_singles Theory.add_axioms_i; |
|
468 |
val add_axiomss = add_multis Theory.add_axioms; |
|
469 |
val add_axiomss_i = add_multis Theory.add_axioms_i; |
|
| 9318 | 470 |
val add_defs = add_singles o Theory.add_defs; |
471 |
val add_defs_i = add_singles o Theory.add_defs_i; |
|
472 |
val add_defss = add_multis o Theory.add_defs; |
|
473 |
val add_defss_i = add_multis o Theory.add_defs_i; |
|
| 4853 | 474 |
end; |
|
4022
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
wenzelm
parents:
4013
diff
changeset
|
475 |
|
|
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
wenzelm
parents:
4013
diff
changeset
|
476 |
|
| 3987 | 477 |
|
|
4963
38aa2d56e28c
added get_name, put_name, global_path, local_path, begin_theory,
wenzelm
parents:
4933
diff
changeset
|
478 |
(*** derived theory operations ***) |
|
38aa2d56e28c
added get_name, put_name, global_path, local_path, begin_theory,
wenzelm
parents:
4933
diff
changeset
|
479 |
|
|
38aa2d56e28c
added get_name, put_name, global_path, local_path, begin_theory,
wenzelm
parents:
4933
diff
changeset
|
480 |
(** theory management **) |
|
38aa2d56e28c
added get_name, put_name, global_path, local_path, begin_theory,
wenzelm
parents:
4933
diff
changeset
|
481 |
|
| 5005 | 482 |
(* data kind 'Pure/theory_management' *) |
|
4963
38aa2d56e28c
added get_name, put_name, global_path, local_path, begin_theory,
wenzelm
parents:
4933
diff
changeset
|
483 |
|
| 5005 | 484 |
structure TheoryManagementDataArgs = |
485 |
struct |
|
486 |
val name = "Pure/theory_management"; |
|
| 6660 | 487 |
type T = {name: string, version: int};
|
| 5000 | 488 |
|
| 6660 | 489 |
val empty = {name = "", version = 0};
|
| 6547 | 490 |
val copy = I; |
| 5005 | 491 |
val prep_ext = I; |
| 5000 | 492 |
fun merge _ = empty; |
| 5005 | 493 |
fun print _ _ = (); |
|
4963
38aa2d56e28c
added get_name, put_name, global_path, local_path, begin_theory,
wenzelm
parents:
4933
diff
changeset
|
494 |
end; |
|
38aa2d56e28c
added get_name, put_name, global_path, local_path, begin_theory,
wenzelm
parents:
4933
diff
changeset
|
495 |
|
| 5005 | 496 |
structure TheoryManagementData = TheoryDataFun(TheoryManagementDataArgs); |
497 |
val get_info = TheoryManagementData.get; |
|
498 |
val put_info = TheoryManagementData.put; |
|
499 |
||
|
4963
38aa2d56e28c
added get_name, put_name, global_path, local_path, begin_theory,
wenzelm
parents:
4933
diff
changeset
|
500 |
|
|
38aa2d56e28c
added get_name, put_name, global_path, local_path, begin_theory,
wenzelm
parents:
4933
diff
changeset
|
501 |
(* get / put name *) |
|
38aa2d56e28c
added get_name, put_name, global_path, local_path, begin_theory,
wenzelm
parents:
4933
diff
changeset
|
502 |
|
| 5000 | 503 |
val get_name = #name o get_info; |
| 6660 | 504 |
fun put_name name = put_info {name = name, version = 0};
|
|
4963
38aa2d56e28c
added get_name, put_name, global_path, local_path, begin_theory,
wenzelm
parents:
4933
diff
changeset
|
505 |
|
|
38aa2d56e28c
added get_name, put_name, global_path, local_path, begin_theory,
wenzelm
parents:
4933
diff
changeset
|
506 |
|
|
38aa2d56e28c
added get_name, put_name, global_path, local_path, begin_theory,
wenzelm
parents:
4933
diff
changeset
|
507 |
(* control prefixing of theory name *) |
|
38aa2d56e28c
added get_name, put_name, global_path, local_path, begin_theory,
wenzelm
parents:
4933
diff
changeset
|
508 |
|
| 5210 | 509 |
val global_path = Theory.root_path; |
|
4963
38aa2d56e28c
added get_name, put_name, global_path, local_path, begin_theory,
wenzelm
parents:
4933
diff
changeset
|
510 |
|
|
38aa2d56e28c
added get_name, put_name, global_path, local_path, begin_theory,
wenzelm
parents:
4933
diff
changeset
|
511 |
fun local_path thy = |
| 5210 | 512 |
thy |> Theory.root_path |> Theory.add_path (get_name thy); |
|
4963
38aa2d56e28c
added get_name, put_name, global_path, local_path, begin_theory,
wenzelm
parents:
4933
diff
changeset
|
513 |
|
|
38aa2d56e28c
added get_name, put_name, global_path, local_path, begin_theory,
wenzelm
parents:
4933
diff
changeset
|
514 |
|
|
38aa2d56e28c
added get_name, put_name, global_path, local_path, begin_theory,
wenzelm
parents:
4933
diff
changeset
|
515 |
(* begin / end theory *) |
|
38aa2d56e28c
added get_name, put_name, global_path, local_path, begin_theory,
wenzelm
parents:
4933
diff
changeset
|
516 |
|
|
38aa2d56e28c
added get_name, put_name, global_path, local_path, begin_theory,
wenzelm
parents:
4933
diff
changeset
|
517 |
fun begin_theory name thys = |
|
38aa2d56e28c
added get_name, put_name, global_path, local_path, begin_theory,
wenzelm
parents:
4933
diff
changeset
|
518 |
Theory.prep_ext_merge thys |
|
38aa2d56e28c
added get_name, put_name, global_path, local_path, begin_theory,
wenzelm
parents:
4933
diff
changeset
|
519 |
|> put_name name |
|
38aa2d56e28c
added get_name, put_name, global_path, local_path, begin_theory,
wenzelm
parents:
4933
diff
changeset
|
520 |
|> local_path; |
|
38aa2d56e28c
added get_name, put_name, global_path, local_path, begin_theory,
wenzelm
parents:
4933
diff
changeset
|
521 |
|
| 12123 | 522 |
fun end_theory thy = |
523 |
thy |
|
524 |
|> Theory.add_name (get_name thy); |
|
|
4963
38aa2d56e28c
added get_name, put_name, global_path, local_path, begin_theory,
wenzelm
parents:
4933
diff
changeset
|
525 |
|
| 6682 | 526 |
fun checkpoint thy = |
527 |
if is_draft thy then |
|
528 |
let val {name, version} = get_info thy in
|
|
529 |
thy |
|
530 |
|> Theory.add_name (name ^ ":" ^ string_of_int version) |
|
531 |
|> put_info {name = name, version = version + 1}
|
|
532 |
end |
|
533 |
else thy; |
|
| 5000 | 534 |
|
535 |
||
|
4963
38aa2d56e28c
added get_name, put_name, global_path, local_path, begin_theory,
wenzelm
parents:
4933
diff
changeset
|
536 |
|
|
38aa2d56e28c
added get_name, put_name, global_path, local_path, begin_theory,
wenzelm
parents:
4933
diff
changeset
|
537 |
(** add logical types **) |
|
4922
03b81b6e1baa
added thms_closure: theory -> xstring -> tthm list option;
wenzelm
parents:
4853
diff
changeset
|
538 |
|
|
03b81b6e1baa
added thms_closure: theory -> xstring -> tthm list option;
wenzelm
parents:
4853
diff
changeset
|
539 |
fun add_typedecls decls thy = |
|
03b81b6e1baa
added thms_closure: theory -> xstring -> tthm list option;
wenzelm
parents:
4853
diff
changeset
|
540 |
let |
|
03b81b6e1baa
added thms_closure: theory -> xstring -> tthm list option;
wenzelm
parents:
4853
diff
changeset
|
541 |
val full = Sign.full_name (Theory.sign_of thy); |
|
03b81b6e1baa
added thms_closure: theory -> xstring -> tthm list option;
wenzelm
parents:
4853
diff
changeset
|
542 |
|
|
03b81b6e1baa
added thms_closure: theory -> xstring -> tthm list option;
wenzelm
parents:
4853
diff
changeset
|
543 |
fun type_of (raw_name, vs, mx) = |
|
03b81b6e1baa
added thms_closure: theory -> xstring -> tthm list option;
wenzelm
parents:
4853
diff
changeset
|
544 |
if null (duplicates vs) then (raw_name, length vs, mx) |
|
03b81b6e1baa
added thms_closure: theory -> xstring -> tthm list option;
wenzelm
parents:
4853
diff
changeset
|
545 |
else error ("Duplicate parameters in type declaration: " ^ quote raw_name);
|
| 14854 | 546 |
in thy |> Theory.add_types (map type_of decls) end; |
|
4922
03b81b6e1baa
added thms_closure: theory -> xstring -> tthm list option;
wenzelm
parents:
4853
diff
changeset
|
547 |
|
|
03b81b6e1baa
added thms_closure: theory -> xstring -> tthm list option;
wenzelm
parents:
4853
diff
changeset
|
548 |
|
|
03b81b6e1baa
added thms_closure: theory -> xstring -> tthm list option;
wenzelm
parents:
4853
diff
changeset
|
549 |
|
| 5091 | 550 |
(*** the ProtoPure theory ***) |
| 3987 | 551 |
|
| 14669 | 552 |
|
553 |
(*It might make sense to restrict the polymorphism of the constant "==" to |
|
554 |
sort logic, instead of the universal sort, {}. Unfortunately, this change
|
|
555 |
causes HOL/Import/shuffler.ML to fail.*) |
|
556 |
||
| 3987 | 557 |
val proto_pure = |
558 |
Theory.pre_pure |
|
|
11516
a0633bdcd015
Added equality axioms and initialization of proof term package.
berghofe
parents:
10667
diff
changeset
|
559 |
|> Library.apply [TheoremsData.init, TheoryManagementData.init, Proofterm.init] |
|
4963
38aa2d56e28c
added get_name, put_name, global_path, local_path, begin_theory,
wenzelm
parents:
4933
diff
changeset
|
560 |
|> put_name "ProtoPure" |
|
38aa2d56e28c
added get_name, put_name, global_path, local_path, begin_theory,
wenzelm
parents:
4933
diff
changeset
|
561 |
|> global_path |
| 3987 | 562 |
|> Theory.add_types |
|
4922
03b81b6e1baa
added thms_closure: theory -> xstring -> tthm list option;
wenzelm
parents:
4853
diff
changeset
|
563 |
[("fun", 2, NoSyn),
|
|
03b81b6e1baa
added thms_closure: theory -> xstring -> tthm list option;
wenzelm
parents:
4853
diff
changeset
|
564 |
("prop", 0, NoSyn),
|
|
03b81b6e1baa
added thms_closure: theory -> xstring -> tthm list option;
wenzelm
parents:
4853
diff
changeset
|
565 |
("itself", 1, NoSyn),
|
|
03b81b6e1baa
added thms_closure: theory -> xstring -> tthm list option;
wenzelm
parents:
4853
diff
changeset
|
566 |
("dummy", 0, NoSyn)]
|
|
03b81b6e1baa
added thms_closure: theory -> xstring -> tthm list option;
wenzelm
parents:
4853
diff
changeset
|
567 |
|> Theory.add_nonterminals Syntax.pure_nonterms |
| 3987 | 568 |
|> Theory.add_syntax Syntax.pure_syntax |
| 6692 | 569 |
|> Theory.add_modesyntax (Symbol.xsymbolsN, true) Syntax.pure_xsym_syntax |
| 3987 | 570 |
|> Theory.add_syntax |
| 7949 | 571 |
[("==>", "[prop, prop] => prop", Delimfix "op ==>"),
|
| 9534 | 572 |
(Term.dummy_patternN, "aprop", Delimfix "'_")] |
| 3987 | 573 |
|> Theory.add_consts |
| 14854 | 574 |
[("==", "['a, 'a] => prop", InfixrName ("==", 2)),
|
| 3987 | 575 |
("==>", "[prop, prop] => prop", Mixfix ("(_/ ==> _)", [2, 1], 1)),
|
576 |
("all", "('a => prop) => prop", Binder ("!!", 0, 0)),
|
|
| 10667 | 577 |
("Goal", "prop => prop", NoSyn),
|
| 6547 | 578 |
("TYPE", "'a itself", NoSyn),
|
| 9534 | 579 |
(Term.dummy_patternN, "'a", Delimfix "'_")] |
|
14223
0ee05eef881b
Added support for making constants final, that is, ensuring that no
skalberg
parents:
13800
diff
changeset
|
580 |
|> Theory.add_finals_i false |
| 14854 | 581 |
[Const("==", [TFree ("'a", []), TFree ("'a", [])] ---> propT),
|
582 |
Const("==>", [propT, propT] ---> propT),
|
|
583 |
Const("all", (TFree("'a", []) --> propT) --> propT),
|
|
584 |
Const("TYPE", a_itselfT)]
|
|
|
5041
a1d0a6d555cd
Goals may now contain assumptions, which are not returned.
nipkow
parents:
5026
diff
changeset
|
585 |
|> Theory.add_modesyntax ("", false)
|
|
12138
7cad58fbc866
renamed open_smart_store_thms to smart_store_thms_open;
wenzelm
parents:
12123
diff
changeset
|
586 |
(Syntax.pure_syntax_output @ Syntax.pure_appl_syntax) |
| 12250 | 587 |
|> Theory.add_trfuns Syntax.pure_trfuns |
588 |
|> Theory.add_trfunsT Syntax.pure_trfunsT |
|
|
4963
38aa2d56e28c
added get_name, put_name, global_path, local_path, begin_theory,
wenzelm
parents:
4933
diff
changeset
|
589 |
|> local_path |
| 10667 | 590 |
|> (#1 oo (add_defs_i false o map Thm.no_attributes)) |
591 |
[("Goal_def", let val A = Free ("A", propT) in Logic.mk_equals (Logic.mk_goal A, A) end)]
|
|
| 9238 | 592 |
|> (#1 o add_thmss [(("nothing", []), [])])
|
|
11516
a0633bdcd015
Added equality axioms and initialization of proof term package.
berghofe
parents:
10667
diff
changeset
|
593 |
|> Theory.add_axioms_i Proofterm.equality_axms |
|
4963
38aa2d56e28c
added get_name, put_name, global_path, local_path, begin_theory,
wenzelm
parents:
4933
diff
changeset
|
594 |
|> end_theory; |
| 3987 | 595 |
|
| 5091 | 596 |
structure ProtoPure = |
597 |
struct |
|
598 |
val thy = proto_pure; |
|
599 |
val Goal_def = get_axiom thy "Goal_def"; |
|
600 |
end; |
|
| 3987 | 601 |
|
602 |
||
603 |
end; |
|
604 |
||
605 |
||
|
4022
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
wenzelm
parents:
4013
diff
changeset
|
606 |
structure BasicPureThy: BASIC_PURE_THY = PureThy; |
|
0770a19c48d3
added ignored_consts, thms_containing, add_store_axioms(_i),
wenzelm
parents:
4013
diff
changeset
|
607 |
open BasicPureThy; |