author | wenzelm |
Mon, 25 Jun 2007 00:36:40 +0200 | |
changeset 23491 | c13ca04303de |
parent 23170 | 94e9413bd7fc |
child 23599 | d889725b0d8a |
permissions | -rw-r--r-- |
22362
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
1 |
(* Title: Pure/more_thm.ML |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
2 |
ID: $Id$ |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
3 |
Author: Makarius |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
4 |
|
22907 | 5 |
Further operations on type ctyp/cterm/thm, outside the inference kernel. |
22362
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
6 |
*) |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
7 |
|
23169 | 8 |
infix aconvc; |
9 |
||
22362
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
10 |
signature THM = |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
11 |
sig |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
12 |
include THM |
23491 | 13 |
val aconvc : cterm * cterm -> bool |
14 |
val add_cterm_frees: cterm -> cterm list -> cterm list |
|
22907 | 15 |
val mk_binop: cterm -> cterm -> cterm -> cterm |
16 |
val dest_binop: cterm -> cterm * cterm |
|
17 |
val dest_implies: cterm -> cterm * cterm |
|
18 |
val dest_equals: cterm -> cterm * cterm |
|
19 |
val dest_equals_lhs: cterm -> cterm |
|
20 |
val dest_equals_rhs: cterm -> cterm |
|
21 |
val lhs_of: thm -> cterm |
|
22 |
val rhs_of: thm -> cterm |
|
22362
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
23 |
val thm_ord: thm * thm -> order |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
24 |
val eq_thm: thm * thm -> bool |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
25 |
val eq_thms: thm list * thm list -> bool |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
26 |
val eq_thm_thy: thm * thm -> bool |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
27 |
val eq_thm_prop: thm * thm -> bool |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
28 |
val equiv_thm: thm * thm -> bool |
22695
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
wenzelm
parents:
22682
diff
changeset
|
29 |
val plain_prop_of: thm -> term |
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
wenzelm
parents:
22682
diff
changeset
|
30 |
val fold_terms: (term -> 'a -> 'a) -> thm -> 'a -> 'a |
22362
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
31 |
val axiomK: string |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
32 |
val assumptionK: string |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
33 |
val definitionK: string |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
34 |
val theoremK: string |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
35 |
val lemmaK: string |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
36 |
val corollaryK: string |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
37 |
val internalK: string |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
38 |
val rule_attribute: (Context.generic -> thm -> thm) -> attribute |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
39 |
val declaration_attribute: (thm -> Context.generic -> Context.generic) -> attribute |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
40 |
val theory_attributes: attribute list -> theory * thm -> theory * thm |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
41 |
val proof_attributes: attribute list -> Proof.context * thm -> Proof.context * thm |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
42 |
val no_attributes: 'a -> 'a * 'b list |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
43 |
val simple_fact: 'a -> ('a * 'b list) list |
22682 | 44 |
val read_def_cterms: |
45 |
theory * (indexname -> typ option) * (indexname -> sort option) -> |
|
46 |
string list -> bool -> (string * typ)list |
|
47 |
-> cterm list * (indexname * typ)list |
|
48 |
val read_cterm: theory -> string * typ -> cterm |
|
22362
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
49 |
end; |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
50 |
|
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
51 |
structure Thm: THM = |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
52 |
struct |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
53 |
|
22695
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
wenzelm
parents:
22682
diff
changeset
|
54 |
(** basic operations **) |
22362
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
55 |
|
23491 | 56 |
(* collecting cterms *) |
57 |
||
58 |
val op aconvc = op aconv o pairself Thm.term_of; |
|
59 |
||
60 |
fun add_cterm_frees ct = |
|
61 |
let |
|
62 |
val cert = Thm.cterm_of (Thm.theory_of_cterm ct); |
|
63 |
val t = Thm.term_of ct; |
|
64 |
in Term.fold_aterms (fn v as Free _ => insert (op aconvc) (cert v) | _ => I) t end; |
|
65 |
||
66 |
||
22907 | 67 |
(* cterm constructors and destructors *) |
68 |
||
69 |
fun mk_binop c a b = Thm.capply (Thm.capply c a) b; |
|
70 |
fun dest_binop ct = (Thm.dest_arg1 ct, Thm.dest_arg ct); |
|
71 |
||
72 |
fun dest_implies ct = |
|
73 |
(case Thm.term_of ct of |
|
74 |
Const ("==>", _) $ _ $ _ => dest_binop ct |
|
75 |
| _ => raise TERM ("dest_implies", [Thm.term_of ct])); |
|
76 |
||
77 |
fun dest_equals ct = |
|
78 |
(case Thm.term_of ct of |
|
79 |
Const ("==", _) $ _ $ _ => dest_binop ct |
|
80 |
| _ => raise TERM ("dest_equals", [Thm.term_of ct])); |
|
81 |
||
82 |
fun dest_equals_lhs ct = |
|
83 |
(case Thm.term_of ct of |
|
84 |
Const ("==", _) $ _ $ _ => Thm.dest_arg1 ct |
|
85 |
| _ => raise TERM ("dest_equals_lhs", [Thm.term_of ct])); |
|
86 |
||
87 |
fun dest_equals_rhs ct = |
|
88 |
(case Thm.term_of ct of |
|
89 |
Const ("==", _) $ _ $ _ => Thm.dest_arg ct |
|
90 |
| _ => raise TERM ("dest_equals_rhs", [Thm.term_of ct])); |
|
91 |
||
92 |
val lhs_of = dest_equals_lhs o Thm.cprop_of; |
|
93 |
val rhs_of = dest_equals_rhs o Thm.cprop_of; |
|
94 |
||
95 |
||
96 |
(* thm order: ignores theory context! *) |
|
22682 | 97 |
|
22362
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
98 |
fun thm_ord (th1, th2) = |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
99 |
let |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
100 |
val {shyps = shyps1, hyps = hyps1, tpairs = tpairs1, prop = prop1, ...} = Thm.rep_thm th1; |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
101 |
val {shyps = shyps2, hyps = hyps2, tpairs = tpairs2, prop = prop2, ...} = Thm.rep_thm th2; |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
102 |
in |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
103 |
(case Term.fast_term_ord (prop1, prop2) of |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
104 |
EQUAL => |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
105 |
(case list_ord (prod_ord Term.fast_term_ord Term.fast_term_ord) (tpairs1, tpairs2) of |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
106 |
EQUAL => |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
107 |
(case list_ord Term.fast_term_ord (hyps1, hyps2) of |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
108 |
EQUAL => list_ord Term.sort_ord (shyps1, shyps2) |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
109 |
| ord => ord) |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
110 |
| ord => ord) |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
111 |
| ord => ord) |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
112 |
end; |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
113 |
|
22682 | 114 |
|
115 |
(* equality *) |
|
116 |
||
22362
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
117 |
fun eq_thm ths = |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
118 |
Context.joinable (pairself Thm.theory_of_thm ths) andalso |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
119 |
thm_ord ths = EQUAL; |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
120 |
|
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
121 |
val eq_thms = eq_list eq_thm; |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
122 |
|
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
123 |
val eq_thm_thy = eq_thy o pairself Thm.theory_of_thm; |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
124 |
val eq_thm_prop = op aconv o pairself Thm.full_prop_of; |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
125 |
|
22682 | 126 |
|
127 |
(* pattern equivalence *) |
|
128 |
||
22362
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
129 |
fun equiv_thm ths = |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
130 |
Pattern.equiv (Theory.merge (pairself Thm.theory_of_thm ths)) (pairself Thm.full_prop_of ths); |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
131 |
|
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
132 |
|
22695
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
wenzelm
parents:
22682
diff
changeset
|
133 |
(* misc operations *) |
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
wenzelm
parents:
22682
diff
changeset
|
134 |
|
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
wenzelm
parents:
22682
diff
changeset
|
135 |
fun plain_prop_of raw_thm = |
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
wenzelm
parents:
22682
diff
changeset
|
136 |
let |
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
wenzelm
parents:
22682
diff
changeset
|
137 |
val thm = Thm.strip_shyps raw_thm; |
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
wenzelm
parents:
22682
diff
changeset
|
138 |
fun err msg = raise THM ("plain_prop_of: " ^ msg, 0, [thm]); |
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
wenzelm
parents:
22682
diff
changeset
|
139 |
val {hyps, prop, tpairs, ...} = Thm.rep_thm thm; |
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
wenzelm
parents:
22682
diff
changeset
|
140 |
in |
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
wenzelm
parents:
22682
diff
changeset
|
141 |
if not (null hyps) then |
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
wenzelm
parents:
22682
diff
changeset
|
142 |
err "theorem may not contain hypotheses" |
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
wenzelm
parents:
22682
diff
changeset
|
143 |
else if not (null (Thm.extra_shyps thm)) then |
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
wenzelm
parents:
22682
diff
changeset
|
144 |
err "theorem may not contain sort hypotheses" |
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
wenzelm
parents:
22682
diff
changeset
|
145 |
else if not (null tpairs) then |
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
wenzelm
parents:
22682
diff
changeset
|
146 |
err "theorem may not contain flex-flex pairs" |
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
wenzelm
parents:
22682
diff
changeset
|
147 |
else prop |
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
wenzelm
parents:
22682
diff
changeset
|
148 |
end; |
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
wenzelm
parents:
22682
diff
changeset
|
149 |
|
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
wenzelm
parents:
22682
diff
changeset
|
150 |
fun fold_terms f th = |
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
wenzelm
parents:
22682
diff
changeset
|
151 |
let val {tpairs, prop, hyps, ...} = Thm.rep_thm th |
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
wenzelm
parents:
22682
diff
changeset
|
152 |
in fold (fn (t, u) => f t #> f u) tpairs #> f prop #> fold f hyps end; |
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
wenzelm
parents:
22682
diff
changeset
|
153 |
|
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
wenzelm
parents:
22682
diff
changeset
|
154 |
|
22682 | 155 |
|
156 |
(** theorem kinds **) |
|
22362
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
157 |
|
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
158 |
val axiomK = "axiom"; |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
159 |
val assumptionK = "assumption"; |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
160 |
val definitionK = "definition"; |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
161 |
val theoremK = "theorem"; |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
162 |
val lemmaK = "lemma"; |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
163 |
val corollaryK = "corollary"; |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
164 |
val internalK = "internal"; |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
165 |
|
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
166 |
|
22682 | 167 |
|
168 |
(** attributes **) |
|
22362
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
169 |
|
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
170 |
fun rule_attribute f (x, th) = (x, f x th); |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
171 |
fun declaration_attribute f (x, th) = (f th x, th); |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
172 |
|
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
173 |
fun apply_attributes mk dest = |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
174 |
let |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
175 |
fun app [] = I |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
176 |
| app ((f: attribute) :: fs) = fn (x, th) => f (mk x, th) |>> dest |> app fs; |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
177 |
in app end; |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
178 |
|
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
179 |
val theory_attributes = apply_attributes Context.Theory Context.the_theory; |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
180 |
val proof_attributes = apply_attributes Context.Proof Context.the_proof; |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
181 |
|
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
182 |
fun no_attributes x = (x, []); |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
183 |
fun simple_fact x = [(x, [])]; |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
184 |
|
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
185 |
|
22682 | 186 |
(** read/certify terms (obsolete) **) (*exception ERROR*) |
187 |
||
188 |
fun read_def_cterms (thy, types, sorts) used freeze sTs = |
|
189 |
let |
|
190 |
val (ts', tye) = Sign.read_def_terms (thy, types, sorts) used freeze sTs; |
|
191 |
val cts = map (Thm.cterm_of thy) ts' |
|
192 |
handle TYPE (msg, _, _) => error msg |
|
193 |
| TERM (msg, _) => error msg; |
|
194 |
in (cts, tye) end; |
|
195 |
||
196 |
fun read_cterm thy sT = |
|
197 |
let val ([ct], _) = read_def_cterms (thy, K NONE, K NONE) [] true [sT] |
|
198 |
in ct end; |
|
199 |
||
200 |
||
22362
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
201 |
open Thm; |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
202 |
|
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
203 |
end; |
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
204 |
|
23170 | 205 |
val op aconvc = Thm.aconvc; |
206 |
||
22362
6470ce514b6e
Further operations on type thm, outside the inference kernel.
wenzelm
parents:
diff
changeset
|
207 |
structure Thmtab = TableFun(type key = thm val ord = Thm.thm_ord); |