author | clasohm |
Tue, 13 Feb 1996 12:57:24 +0100 | |
changeset 1493 | e936723cb94d |
parent 1460 | 5a6f2aabd538 |
child 1495 | b8b54847c77f |
permissions | -rw-r--r-- |
250 | 1 |
(* Title: Pure/thm.ML |
0 | 2 |
ID: $Id$ |
250 | 3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
229
4002c4cd450c
Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents:
225
diff
changeset
|
4 |
Copyright 1994 University of Cambridge |
4002c4cd450c
Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents:
225
diff
changeset
|
5 |
|
1160 | 6 |
The core of Isabelle's Meta Logic: certified types and terms, meta |
7 |
theorems, theories, meta rules (including resolution and |
|
8 |
simplification). |
|
0 | 9 |
*) |
10 |
||
250 | 11 |
signature THM = |
12 |
sig |
|
1238 | 13 |
structure Envir : ENVIR |
14 |
structure Sequence : SEQUENCE |
|
15 |
structure Sign : SIGN |
|
1160 | 16 |
|
17 |
(*certified types*) |
|
387
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
18 |
type ctyp |
1238 | 19 |
val rep_ctyp : ctyp -> {sign: Sign.sg, T: typ} |
20 |
val typ_of : ctyp -> typ |
|
21 |
val ctyp_of : Sign.sg -> typ -> ctyp |
|
22 |
val read_ctyp : Sign.sg -> string -> ctyp |
|
1160 | 23 |
|
24 |
(*certified terms*) |
|
25 |
type cterm |
|
1493
e936723cb94d
added dest_comb, dest_abs and mk_prop for manipulating cterms
clasohm
parents:
1460
diff
changeset
|
26 |
exception CTERM of string |
e936723cb94d
added dest_comb, dest_abs and mk_prop for manipulating cterms
clasohm
parents:
1460
diff
changeset
|
27 |
val rep_cterm : cterm -> {sign: Sign.sg, t: term, T: typ, |
e936723cb94d
added dest_comb, dest_abs and mk_prop for manipulating cterms
clasohm
parents:
1460
diff
changeset
|
28 |
maxidx: int} |
1238 | 29 |
val term_of : cterm -> term |
30 |
val cterm_of : Sign.sg -> term -> cterm |
|
31 |
val read_cterm : Sign.sg -> string * typ -> cterm |
|
1394
a1d2735f5ade
New function read_cterms is a combination of read_def_cterm and
paulson
parents:
1258
diff
changeset
|
32 |
val read_cterms : Sign.sg -> string list * typ list -> cterm list |
1238 | 33 |
val cterm_fun : (term -> term) -> (cterm -> cterm) |
34 |
val dest_cimplies : cterm -> cterm * cterm |
|
1493
e936723cb94d
added dest_comb, dest_abs and mk_prop for manipulating cterms
clasohm
parents:
1460
diff
changeset
|
35 |
val dest_comb : cterm -> cterm * cterm |
e936723cb94d
added dest_comb, dest_abs and mk_prop for manipulating cterms
clasohm
parents:
1460
diff
changeset
|
36 |
val dest_abs : cterm -> cterm * cterm |
e936723cb94d
added dest_comb, dest_abs and mk_prop for manipulating cterms
clasohm
parents:
1460
diff
changeset
|
37 |
val mk_prop : cterm -> cterm |
1238 | 38 |
val read_def_cterm : |
1160 | 39 |
Sign.sg * (indexname -> typ option) * (indexname -> sort option) -> |
40 |
string list -> bool -> string * typ -> cterm * (indexname * typ) list |
|
41 |
||
42 |
(*meta theorems*) |
|
43 |
type thm |
|
44 |
exception THM of string * int * thm list |
|
1238 | 45 |
val rep_thm : thm -> {sign: Sign.sg, maxidx: int, |
1220 | 46 |
shyps: sort list, hyps: term list, prop: term} |
1238 | 47 |
val stamps_of_thm : thm -> string ref list |
48 |
val tpairs_of : thm -> (term * term) list |
|
49 |
val prems_of : thm -> term list |
|
50 |
val nprems_of : thm -> int |
|
51 |
val concl_of : thm -> term |
|
52 |
val cprop_of : thm -> cterm |
|
53 |
val extra_shyps : thm -> sort list |
|
54 |
val force_strip_shyps : bool ref (* FIXME tmp *) |
|
55 |
val strip_shyps : thm -> thm |
|
56 |
val implies_intr_shyps: thm -> thm |
|
57 |
val cert_axm : Sign.sg -> string * term -> string * term |
|
58 |
val read_axm : Sign.sg -> string * string -> string * term |
|
59 |
val inferT_axm : Sign.sg -> string * term -> string * term |
|
1160 | 60 |
|
61 |
(*theories*) |
|
62 |
type theory |
|
63 |
exception THEORY of string * theory list |
|
1238 | 64 |
val rep_theory : theory -> |
1160 | 65 |
{sign: Sign.sg, new_axioms: term Sign.Symtab.table, parents: theory list} |
1238 | 66 |
val sign_of : theory -> Sign.sg |
67 |
val syn_of : theory -> Sign.Syntax.syntax |
|
68 |
val stamps_of_thy : theory -> string ref list |
|
69 |
val parents_of : theory -> theory list |
|
70 |
val subthy : theory * theory -> bool |
|
71 |
val eq_thy : theory * theory -> bool |
|
72 |
val get_axiom : theory -> string -> thm |
|
73 |
val axioms_of : theory -> (string * thm) list |
|
74 |
val proto_pure_thy : theory |
|
75 |
val pure_thy : theory |
|
76 |
val cpure_thy : theory |
|
564 | 77 |
local open Sign.Syntax in |
1238 | 78 |
val add_classes : (class * class list) list -> theory -> theory |
79 |
val add_classrel : (class * class) list -> theory -> theory |
|
80 |
val add_defsort : sort -> theory -> theory |
|
81 |
val add_types : (string * int * mixfix) list -> theory -> theory |
|
82 |
val add_tyabbrs : (string * string list * string * mixfix) list |
|
387
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
83 |
-> theory -> theory |
1238 | 84 |
val add_tyabbrs_i : (string * string list * typ * mixfix) list |
387
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
85 |
-> theory -> theory |
1238 | 86 |
val add_arities : (string * sort list * sort) list -> theory -> theory |
87 |
val add_consts : (string * string * mixfix) list -> theory -> theory |
|
88 |
val add_consts_i : (string * typ * mixfix) list -> theory -> theory |
|
89 |
val add_syntax : (string * string * mixfix) list -> theory -> theory |
|
90 |
val add_syntax_i : (string * typ * mixfix) list -> theory -> theory |
|
91 |
val add_trfuns : |
|
387
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
92 |
(string * (ast list -> ast)) list * |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
93 |
(string * (term list -> term)) list * |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
94 |
(string * (term list -> term)) list * |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
95 |
(string * (ast list -> ast)) list -> theory -> theory |
1238 | 96 |
val add_trrules : (string * string) trrule list -> theory -> theory |
97 |
val add_trrules_i : ast trrule list -> theory -> theory |
|
98 |
val add_axioms : (string * string) list -> theory -> theory |
|
99 |
val add_axioms_i : (string * term) list -> theory -> theory |
|
100 |
val add_thyname : string -> theory -> theory |
|
387
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
101 |
end |
1238 | 102 |
val merge_theories : theory * theory -> theory |
103 |
val merge_thy_list : bool -> theory list -> theory |
|
1160 | 104 |
|
105 |
(*meta rules*) |
|
1238 | 106 |
val assume : cterm -> thm |
1416 | 107 |
val compress : thm -> thm |
1238 | 108 |
val implies_intr : cterm -> thm -> thm |
109 |
val implies_elim : thm -> thm -> thm |
|
110 |
val forall_intr : cterm -> thm -> thm |
|
111 |
val forall_elim : cterm -> thm -> thm |
|
112 |
val flexpair_def : thm |
|
113 |
val reflexive : cterm -> thm |
|
114 |
val symmetric : thm -> thm |
|
115 |
val transitive : thm -> thm -> thm |
|
116 |
val beta_conversion : cterm -> thm |
|
117 |
val extensional : thm -> thm |
|
118 |
val abstract_rule : string -> cterm -> thm -> thm |
|
119 |
val combination : thm -> thm -> thm |
|
120 |
val equal_intr : thm -> thm -> thm |
|
121 |
val equal_elim : thm -> thm -> thm |
|
122 |
val implies_intr_hyps : thm -> thm |
|
123 |
val flexflex_rule : thm -> thm Sequence.seq |
|
124 |
val instantiate : |
|
1160 | 125 |
(indexname * ctyp) list * (cterm * cterm) list -> thm -> thm |
1238 | 126 |
val trivial : cterm -> thm |
127 |
val class_triv : theory -> class -> thm |
|
128 |
val varifyT : thm -> thm |
|
129 |
val freezeT : thm -> thm |
|
130 |
val dest_state : thm * int -> |
|
1160 | 131 |
(term * term) list * term list * term * term |
1238 | 132 |
val lift_rule : (thm * int) -> thm -> thm |
133 |
val assumption : int -> thm -> thm Sequence.seq |
|
134 |
val eq_assumption : int -> thm -> thm |
|
1160 | 135 |
val rename_params_rule: string list * int -> thm -> thm |
1238 | 136 |
val bicompose : bool -> bool * thm * int -> |
1160 | 137 |
int -> thm -> thm Sequence.seq |
1238 | 138 |
val biresolution : bool -> (bool * thm) list -> |
1160 | 139 |
int -> thm -> thm Sequence.seq |
140 |
||
141 |
(*meta simplification*) |
|
142 |
type meta_simpset |
|
143 |
exception SIMPLIFIER of string * thm |
|
1238 | 144 |
val empty_mss : meta_simpset |
145 |
val add_simps : meta_simpset * thm list -> meta_simpset |
|
146 |
val del_simps : meta_simpset * thm list -> meta_simpset |
|
147 |
val mss_of : thm list -> meta_simpset |
|
148 |
val add_congs : meta_simpset * thm list -> meta_simpset |
|
149 |
val add_prems : meta_simpset * thm list -> meta_simpset |
|
150 |
val prems_of_mss : meta_simpset -> thm list |
|
151 |
val set_mk_rews : meta_simpset * (thm -> thm list) -> meta_simpset |
|
152 |
val mk_rews_of_mss : meta_simpset -> thm -> thm list |
|
153 |
val trace_simp : bool ref |
|
154 |
val rewrite_cterm : bool * bool -> meta_simpset -> |
|
1160 | 155 |
(meta_simpset -> thm -> thm option) -> cterm -> thm |
250 | 156 |
end; |
0 | 157 |
|
250 | 158 |
functor ThmFun (structure Logic: LOGIC and Unify: UNIFY and Pattern: PATTERN |
399
86cc2b98f9e0
added class_triv: theory -> class -> thm (for axclasses);
wenzelm
parents:
387
diff
changeset
|
159 |
and Net:NET sharing type Pattern.type_sig = Unify.Sign.Type.type_sig): THM = |
0 | 160 |
struct |
250 | 161 |
|
0 | 162 |
structure Sequence = Unify.Sequence; |
163 |
structure Envir = Unify.Envir; |
|
164 |
structure Sign = Unify.Sign; |
|
165 |
structure Type = Sign.Type; |
|
166 |
structure Syntax = Sign.Syntax; |
|
167 |
structure Symtab = Sign.Symtab; |
|
168 |
||
169 |
||
387
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
170 |
(*** Certified terms and types ***) |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
171 |
|
250 | 172 |
(** certified types **) |
173 |
||
174 |
(*certified typs under a signature*) |
|
175 |
||
176 |
datatype ctyp = Ctyp of {sign: Sign.sg, T: typ}; |
|
177 |
||
178 |
fun rep_ctyp (Ctyp args) = args; |
|
179 |
fun typ_of (Ctyp {T, ...}) = T; |
|
180 |
||
181 |
fun ctyp_of sign T = |
|
182 |
Ctyp {sign = sign, T = Sign.certify_typ sign T}; |
|
183 |
||
184 |
fun read_ctyp sign s = |
|
185 |
Ctyp {sign = sign, T = Sign.read_typ (sign, K None) s}; |
|
229
4002c4cd450c
Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents:
225
diff
changeset
|
186 |
|
4002c4cd450c
Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents:
225
diff
changeset
|
187 |
|
4002c4cd450c
Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents:
225
diff
changeset
|
188 |
|
250 | 189 |
(** certified terms **) |
229
4002c4cd450c
Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents:
225
diff
changeset
|
190 |
|
250 | 191 |
(*certified terms under a signature, with checked typ and maxidx of Vars*) |
229
4002c4cd450c
Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents:
225
diff
changeset
|
192 |
|
250 | 193 |
datatype cterm = Cterm of {sign: Sign.sg, t: term, T: typ, maxidx: int}; |
229
4002c4cd450c
Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents:
225
diff
changeset
|
194 |
|
4002c4cd450c
Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents:
225
diff
changeset
|
195 |
fun rep_cterm (Cterm args) = args; |
250 | 196 |
fun term_of (Cterm {t, ...}) = t; |
229
4002c4cd450c
Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents:
225
diff
changeset
|
197 |
|
250 | 198 |
(*create a cterm by checking a "raw" term with respect to a signature*) |
199 |
fun cterm_of sign tm = |
|
200 |
let val (t, T, maxidx) = Sign.certify_term sign tm |
|
1394
a1d2735f5ade
New function read_cterms is a combination of read_def_cterm and
paulson
parents:
1258
diff
changeset
|
201 |
in Cterm {sign = sign, t = t, T = T, maxidx = maxidx} |
a1d2735f5ade
New function read_cterms is a combination of read_def_cterm and
paulson
parents:
1258
diff
changeset
|
202 |
end; |
229
4002c4cd450c
Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents:
225
diff
changeset
|
203 |
|
250 | 204 |
fun cterm_fun f (Cterm {sign, t, ...}) = cterm_of sign (f t); |
205 |
||
229
4002c4cd450c
Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents:
225
diff
changeset
|
206 |
|
250 | 207 |
(*dest_implies for cterms. Note T=prop below*) |
208 |
fun dest_cimplies (Cterm{sign, T, maxidx, t=Const("==>", _) $ A $ B}) = |
|
229
4002c4cd450c
Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents:
225
diff
changeset
|
209 |
(Cterm{sign=sign, T=T, maxidx=maxidx, t=A}, |
250 | 210 |
Cterm{sign=sign, T=T, maxidx=maxidx, t=B}) |
211 |
| dest_cimplies ct = raise TERM ("dest_cimplies", [term_of ct]); |
|
229
4002c4cd450c
Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents:
225
diff
changeset
|
212 |
|
1493
e936723cb94d
added dest_comb, dest_abs and mk_prop for manipulating cterms
clasohm
parents:
1460
diff
changeset
|
213 |
exception CTERM of string; |
e936723cb94d
added dest_comb, dest_abs and mk_prop for manipulating cterms
clasohm
parents:
1460
diff
changeset
|
214 |
|
e936723cb94d
added dest_comb, dest_abs and mk_prop for manipulating cterms
clasohm
parents:
1460
diff
changeset
|
215 |
(*Destruct application in cterms*) |
e936723cb94d
added dest_comb, dest_abs and mk_prop for manipulating cterms
clasohm
parents:
1460
diff
changeset
|
216 |
fun dest_comb (Cterm{sign, T, maxidx, t = A $ B}) = |
e936723cb94d
added dest_comb, dest_abs and mk_prop for manipulating cterms
clasohm
parents:
1460
diff
changeset
|
217 |
let val typeA = fastype_of A; |
e936723cb94d
added dest_comb, dest_abs and mk_prop for manipulating cterms
clasohm
parents:
1460
diff
changeset
|
218 |
val typeB = |
e936723cb94d
added dest_comb, dest_abs and mk_prop for manipulating cterms
clasohm
parents:
1460
diff
changeset
|
219 |
case typeA of Type("fun",[S,T]) => S |
e936723cb94d
added dest_comb, dest_abs and mk_prop for manipulating cterms
clasohm
parents:
1460
diff
changeset
|
220 |
| _ => error "Function type expected in dest_comb"; |
e936723cb94d
added dest_comb, dest_abs and mk_prop for manipulating cterms
clasohm
parents:
1460
diff
changeset
|
221 |
in |
e936723cb94d
added dest_comb, dest_abs and mk_prop for manipulating cterms
clasohm
parents:
1460
diff
changeset
|
222 |
(Cterm {sign=sign, maxidx=maxidx, t=A, T=typeA}, |
e936723cb94d
added dest_comb, dest_abs and mk_prop for manipulating cterms
clasohm
parents:
1460
diff
changeset
|
223 |
Cterm {sign=sign, maxidx=maxidx, t=B, T=typeB}) |
e936723cb94d
added dest_comb, dest_abs and mk_prop for manipulating cterms
clasohm
parents:
1460
diff
changeset
|
224 |
end |
e936723cb94d
added dest_comb, dest_abs and mk_prop for manipulating cterms
clasohm
parents:
1460
diff
changeset
|
225 |
| dest_comb _ = raise CTERM "dest_comb"; |
e936723cb94d
added dest_comb, dest_abs and mk_prop for manipulating cterms
clasohm
parents:
1460
diff
changeset
|
226 |
|
e936723cb94d
added dest_comb, dest_abs and mk_prop for manipulating cterms
clasohm
parents:
1460
diff
changeset
|
227 |
(*Destruct abstraction in cterms*) |
e936723cb94d
added dest_comb, dest_abs and mk_prop for manipulating cterms
clasohm
parents:
1460
diff
changeset
|
228 |
fun dest_abs (Cterm{sign, T, maxidx, t = tm as Abs(s,ty,M)}) = |
e936723cb94d
added dest_comb, dest_abs and mk_prop for manipulating cterms
clasohm
parents:
1460
diff
changeset
|
229 |
let fun mk_var{Name,Ty} = Free(Name,Ty); |
e936723cb94d
added dest_comb, dest_abs and mk_prop for manipulating cterms
clasohm
parents:
1460
diff
changeset
|
230 |
val v = mk_var{Name = s, Ty = ty}; |
e936723cb94d
added dest_comb, dest_abs and mk_prop for manipulating cterms
clasohm
parents:
1460
diff
changeset
|
231 |
val ty2 = |
e936723cb94d
added dest_comb, dest_abs and mk_prop for manipulating cterms
clasohm
parents:
1460
diff
changeset
|
232 |
case T of Type("fun",[_,S]) => S |
e936723cb94d
added dest_comb, dest_abs and mk_prop for manipulating cterms
clasohm
parents:
1460
diff
changeset
|
233 |
| _ => error "Function type expected in dest_abs"; |
e936723cb94d
added dest_comb, dest_abs and mk_prop for manipulating cterms
clasohm
parents:
1460
diff
changeset
|
234 |
in (Cterm{sign = sign, T = ty, maxidx = maxidx, t = v}, |
e936723cb94d
added dest_comb, dest_abs and mk_prop for manipulating cterms
clasohm
parents:
1460
diff
changeset
|
235 |
Cterm{sign = sign, T = ty2, maxidx = maxidx, t = betapply (tm,v)}) |
e936723cb94d
added dest_comb, dest_abs and mk_prop for manipulating cterms
clasohm
parents:
1460
diff
changeset
|
236 |
end |
e936723cb94d
added dest_comb, dest_abs and mk_prop for manipulating cterms
clasohm
parents:
1460
diff
changeset
|
237 |
| dest_abs _ = raise CTERM "dest_abs"; |
e936723cb94d
added dest_comb, dest_abs and mk_prop for manipulating cterms
clasohm
parents:
1460
diff
changeset
|
238 |
|
e936723cb94d
added dest_comb, dest_abs and mk_prop for manipulating cterms
clasohm
parents:
1460
diff
changeset
|
239 |
(*Convert cterm of type "o" to "prop" by using Trueprop*) |
e936723cb94d
added dest_comb, dest_abs and mk_prop for manipulating cterms
clasohm
parents:
1460
diff
changeset
|
240 |
fun mk_prop (ct as Cterm{sign, T, maxidx, t = Const("Trueprop",_) $ _}) = ct |
e936723cb94d
added dest_comb, dest_abs and mk_prop for manipulating cterms
clasohm
parents:
1460
diff
changeset
|
241 |
| mk_prop (Cterm{sign, T, maxidx, t}) = |
e936723cb94d
added dest_comb, dest_abs and mk_prop for manipulating cterms
clasohm
parents:
1460
diff
changeset
|
242 |
if T = Type("o",[]) then |
e936723cb94d
added dest_comb, dest_abs and mk_prop for manipulating cterms
clasohm
parents:
1460
diff
changeset
|
243 |
Cterm{sign = sign, T = Type("prop",[]), maxidx = maxidx, |
e936723cb94d
added dest_comb, dest_abs and mk_prop for manipulating cterms
clasohm
parents:
1460
diff
changeset
|
244 |
t = Const("Trueprop", Type("o",[]) --> Type("prop",[])) $ t} |
e936723cb94d
added dest_comb, dest_abs and mk_prop for manipulating cterms
clasohm
parents:
1460
diff
changeset
|
245 |
else error "Type o expected in mk_prop"; |
250 | 246 |
|
229
4002c4cd450c
Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents:
225
diff
changeset
|
247 |
|
574 | 248 |
(** read cterms **) (*exception ERROR*) |
250 | 249 |
|
250 |
(*read term, infer types, certify term*) |
|
949
83c588d6fee9
Changed treatment of during type inference internally generated type
nipkow
parents:
922
diff
changeset
|
251 |
fun read_def_cterm (sign, types, sorts) used freeze (a, T) = |
250 | 252 |
let |
574 | 253 |
val T' = Sign.certify_typ sign T |
254 |
handle TYPE (msg, _, _) => error msg; |
|
623 | 255 |
val ts = Syntax.read (#syn (Sign.rep_sg sign)) T' a; |
949
83c588d6fee9
Changed treatment of during type inference internally generated type
nipkow
parents:
922
diff
changeset
|
256 |
val (_, t', tye) = |
959 | 257 |
Sign.infer_types sign types sorts used freeze (ts, T'); |
574 | 258 |
val ct = cterm_of sign t' |
1394
a1d2735f5ade
New function read_cterms is a combination of read_def_cterm and
paulson
parents:
1258
diff
changeset
|
259 |
handle TYPE arg => error (Sign.exn_type_msg sign arg) |
1460 | 260 |
| TERM (msg, _) => error msg; |
250 | 261 |
in (ct, tye) end; |
229
4002c4cd450c
Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents:
225
diff
changeset
|
262 |
|
949
83c588d6fee9
Changed treatment of during type inference internally generated type
nipkow
parents:
922
diff
changeset
|
263 |
fun read_cterm sign = #1 o read_def_cterm (sign, K None, K None) [] true; |
229
4002c4cd450c
Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents:
225
diff
changeset
|
264 |
|
1394
a1d2735f5ade
New function read_cterms is a combination of read_def_cterm and
paulson
parents:
1258
diff
changeset
|
265 |
(*read a list of terms, matching them against a list of expected types. |
a1d2735f5ade
New function read_cterms is a combination of read_def_cterm and
paulson
parents:
1258
diff
changeset
|
266 |
NO disambiguation of alternative parses via type-checking -- it is just |
a1d2735f5ade
New function read_cterms is a combination of read_def_cterm and
paulson
parents:
1258
diff
changeset
|
267 |
not practical.*) |
a1d2735f5ade
New function read_cterms is a combination of read_def_cterm and
paulson
parents:
1258
diff
changeset
|
268 |
fun read_cterms sign (bs, Ts) = |
a1d2735f5ade
New function read_cterms is a combination of read_def_cterm and
paulson
parents:
1258
diff
changeset
|
269 |
let |
a1d2735f5ade
New function read_cterms is a combination of read_def_cterm and
paulson
parents:
1258
diff
changeset
|
270 |
val {tsig, syn, ...} = Sign.rep_sg sign |
a1d2735f5ade
New function read_cterms is a combination of read_def_cterm and
paulson
parents:
1258
diff
changeset
|
271 |
fun read (b,T) = |
1460 | 272 |
case Syntax.read syn T b of |
273 |
[t] => t |
|
274 |
| _ => error("Error or ambiguity in parsing of " ^ b) |
|
1394
a1d2735f5ade
New function read_cterms is a combination of read_def_cterm and
paulson
parents:
1258
diff
changeset
|
275 |
val (us,_) = Type.infer_types(tsig, Sign.const_type sign, |
1460 | 276 |
K None, K None, |
277 |
[], true, |
|
278 |
map (Sign.certify_typ sign) Ts, |
|
279 |
map read (bs~~Ts)) |
|
1394
a1d2735f5ade
New function read_cterms is a combination of read_def_cterm and
paulson
parents:
1258
diff
changeset
|
280 |
in map (cterm_of sign) us end |
a1d2735f5ade
New function read_cterms is a combination of read_def_cterm and
paulson
parents:
1258
diff
changeset
|
281 |
handle TYPE arg => error (Sign.exn_type_msg sign arg) |
a1d2735f5ade
New function read_cterms is a combination of read_def_cterm and
paulson
parents:
1258
diff
changeset
|
282 |
| TERM (msg, _) => error msg; |
a1d2735f5ade
New function read_cterms is a combination of read_def_cterm and
paulson
parents:
1258
diff
changeset
|
283 |
|
250 | 284 |
|
285 |
||
387
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
286 |
(*** Meta theorems ***) |
229
4002c4cd450c
Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents:
225
diff
changeset
|
287 |
|
0 | 288 |
datatype thm = Thm of |
1460 | 289 |
{sign: Sign.sg, (*signature for hyps and prop*) |
290 |
maxidx: int, (*maximum index of any Var or TVar*) |
|
291 |
shyps: sort list, (* FIXME comment *) |
|
292 |
hyps: term list, (*hypotheses*) |
|
293 |
prop: term}; (*conclusion*) |
|
0 | 294 |
|
250 | 295 |
fun rep_thm (Thm args) = args; |
0 | 296 |
|
387
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
297 |
(*errors involving theorems*) |
0 | 298 |
exception THM of string * int * thm list; |
299 |
||
387
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
300 |
|
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
301 |
val sign_of_thm = #sign o rep_thm; |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
302 |
val stamps_of_thm = #stamps o Sign.rep_sg o sign_of_thm; |
0 | 303 |
|
387
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
304 |
(*merge signatures of two theorems; raise exception if incompatible*) |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
305 |
fun merge_thm_sgs (th1, th2) = |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
306 |
Sign.merge (pairself sign_of_thm (th1, th2)) |
574 | 307 |
handle TERM (msg, _) => raise THM (msg, 0, [th1, th2]); |
387
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
308 |
|
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
309 |
|
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
310 |
(*maps object-rule to tpairs*) |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
311 |
fun tpairs_of (Thm {prop, ...}) = #1 (Logic.strip_flexpairs prop); |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
312 |
|
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
313 |
(*maps object-rule to premises*) |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
314 |
fun prems_of (Thm {prop, ...}) = |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
315 |
Logic.strip_imp_prems (Logic.skip_flexpairs prop); |
0 | 316 |
|
317 |
(*counts premises in a rule*) |
|
387
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
318 |
fun nprems_of (Thm {prop, ...}) = |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
319 |
Logic.count_prems (Logic.skip_flexpairs prop, 0); |
0 | 320 |
|
387
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
321 |
(*maps object-rule to conclusion*) |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
322 |
fun concl_of (Thm {prop, ...}) = Logic.strip_imp_concl prop; |
0 | 323 |
|
387
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
324 |
(*the statement of any thm is a cterm*) |
1160 | 325 |
fun cprop_of (Thm {sign, maxidx, prop, ...}) = |
387
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
326 |
Cterm {sign = sign, maxidx = maxidx, T = propT, t = prop}; |
229
4002c4cd450c
Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents:
225
diff
changeset
|
327 |
|
387
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
328 |
|
0 | 329 |
|
1238 | 330 |
(** sort contexts of theorems **) |
331 |
||
332 |
(* basic utils *) |
|
333 |
||
334 |
(*accumulate sorts suppressing duplicates; these are coded low level |
|
335 |
to improve efficiency a bit*) |
|
336 |
||
337 |
fun add_typ_sorts (Type (_, Ts), Ss) = add_typs_sorts (Ts, Ss) |
|
338 |
| add_typ_sorts (TFree (_, S), Ss) = S ins Ss |
|
339 |
| add_typ_sorts (TVar (_, S), Ss) = S ins Ss |
|
340 |
and add_typs_sorts ([], Ss) = Ss |
|
341 |
| add_typs_sorts (T :: Ts, Ss) = add_typs_sorts (Ts, add_typ_sorts (T, Ss)); |
|
342 |
||
343 |
fun add_term_sorts (Const (_, T), Ss) = add_typ_sorts (T, Ss) |
|
344 |
| add_term_sorts (Free (_, T), Ss) = add_typ_sorts (T, Ss) |
|
345 |
| add_term_sorts (Var (_, T), Ss) = add_typ_sorts (T, Ss) |
|
346 |
| add_term_sorts (Bound _, Ss) = Ss |
|
347 |
| add_term_sorts (Abs (_, T, t), Ss) = add_term_sorts (t, add_typ_sorts (T, Ss)) |
|
348 |
| add_term_sorts (t $ u, Ss) = add_term_sorts (t, add_term_sorts (u, Ss)); |
|
349 |
||
350 |
fun add_terms_sorts ([], Ss) = Ss |
|
351 |
| add_terms_sorts (t :: ts, Ss) = add_terms_sorts (ts, add_term_sorts (t, Ss)); |
|
352 |
||
1258 | 353 |
fun env_codT (Envir.Envir {iTs, ...}) = map snd iTs; |
354 |
||
355 |
fun add_env_sorts (env, Ss) = |
|
356 |
add_terms_sorts (map snd (Envir.alist_of env), |
|
357 |
add_typs_sorts (env_codT env, Ss)); |
|
358 |
||
1238 | 359 |
fun add_thm_sorts (Thm {hyps, prop, ...}, Ss) = |
360 |
add_terms_sorts (hyps, add_term_sorts (prop, Ss)); |
|
361 |
||
362 |
fun add_thms_shyps ([], Ss) = Ss |
|
363 |
| add_thms_shyps (Thm {shyps, ...} :: ths, Ss) = |
|
364 |
add_thms_shyps (ths, shyps union Ss); |
|
365 |
||
366 |
||
367 |
(*get 'dangling' sort constraints of a thm*) |
|
368 |
fun extra_shyps (th as Thm {shyps, ...}) = |
|
369 |
shyps \\ add_thm_sorts (th, []); |
|
370 |
||
371 |
||
1416 | 372 |
(*Compression of theorems -- a separate rule, not integrated with the others, |
373 |
as it could be slow.*) |
|
374 |
fun compress (Thm {sign, maxidx, shyps, hyps, prop}) = |
|
375 |
Thm {sign = sign, |
|
1460 | 376 |
maxidx = maxidx, |
377 |
shyps = shyps, |
|
378 |
hyps = map Term.compress_term hyps, |
|
379 |
prop = Term.compress_term prop}; |
|
1416 | 380 |
|
381 |
||
1238 | 382 |
(* fix_shyps *) |
383 |
||
384 |
(*preserve sort contexts of rule premises and substituted types*) |
|
385 |
fun fix_shyps thms Ts thm = |
|
386 |
let |
|
387 |
val Thm {sign, maxidx, hyps, prop, ...} = thm; |
|
388 |
val shyps = |
|
389 |
add_thm_sorts (thm, add_typs_sorts (Ts, add_thms_shyps (thms, []))); |
|
390 |
in |
|
391 |
Thm {sign = sign, maxidx = maxidx, |
|
392 |
shyps = shyps, hyps = hyps, prop = prop} |
|
393 |
end; |
|
394 |
||
395 |
||
396 |
(* strip_shyps *) (* FIXME improve? (e.g. only minimal extra sorts) *) |
|
397 |
||
398 |
val force_strip_shyps = ref true; (* FIXME tmp *) |
|
399 |
||
400 |
(*remove extra sorts that are known to be syntactically non-empty*) |
|
401 |
fun strip_shyps thm = |
|
402 |
let |
|
403 |
val Thm {sign, maxidx, shyps, hyps, prop} = thm; |
|
404 |
val sorts = add_thm_sorts (thm, []); |
|
405 |
val maybe_empty = not o Sign.nonempty_sort sign sorts; |
|
406 |
val shyps' = filter (fn S => S mem sorts orelse maybe_empty S) shyps; |
|
407 |
in |
|
408 |
Thm {sign = sign, maxidx = maxidx, |
|
409 |
shyps = |
|
410 |
(if eq_set (shyps', sorts) orelse not (! force_strip_shyps) then shyps' |
|
411 |
else (* FIXME tmp *) |
|
412 |
(writeln ("WARNING Removed sort hypotheses: " ^ |
|
413 |
commas (map Type.str_of_sort (shyps' \\ sorts))); |
|
414 |
writeln "WARNING Let's hope these sorts are non-empty!"; |
|
415 |
sorts)), |
|
416 |
hyps = hyps, prop = prop} |
|
417 |
end; |
|
418 |
||
419 |
||
420 |
(* implies_intr_shyps *) |
|
421 |
||
422 |
(*discharge all extra sort hypotheses*) |
|
423 |
fun implies_intr_shyps thm = |
|
424 |
(case extra_shyps thm of |
|
425 |
[] => thm |
|
426 |
| xshyps => |
|
427 |
let |
|
428 |
val Thm {sign, maxidx, shyps, hyps, prop} = thm; |
|
429 |
val shyps' = logicS ins (shyps \\ xshyps); |
|
430 |
val used_names = foldr add_term_tfree_names (prop :: hyps, []); |
|
431 |
val names = |
|
432 |
tl (variantlist (replicate (length xshyps + 1) "'", used_names)); |
|
433 |
val tfrees = map (TFree o rpair logicS) names; |
|
434 |
||
435 |
fun mk_insort (T, S) = map (Logic.mk_inclass o pair T) S; |
|
436 |
val sort_hyps = flat (map2 mk_insort (tfrees, xshyps)); |
|
437 |
in |
|
438 |
Thm {sign = sign, maxidx = maxidx, shyps = shyps', |
|
439 |
hyps = hyps, prop = Logic.list_implies (sort_hyps, prop)} |
|
440 |
end); |
|
441 |
||
442 |
||
443 |
||
387
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
444 |
(*** Theories ***) |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
445 |
|
0 | 446 |
datatype theory = |
399
86cc2b98f9e0
added class_triv: theory -> class -> thm (for axclasses);
wenzelm
parents:
387
diff
changeset
|
447 |
Theory of { |
86cc2b98f9e0
added class_triv: theory -> class -> thm (for axclasses);
wenzelm
parents:
387
diff
changeset
|
448 |
sign: Sign.sg, |
86cc2b98f9e0
added class_triv: theory -> class -> thm (for axclasses);
wenzelm
parents:
387
diff
changeset
|
449 |
new_axioms: term Symtab.table, |
86cc2b98f9e0
added class_triv: theory -> class -> thm (for axclasses);
wenzelm
parents:
387
diff
changeset
|
450 |
parents: theory list}; |
0 | 451 |
|
387
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
452 |
fun rep_theory (Theory args) = args; |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
453 |
|
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
454 |
(*errors involving theories*) |
0 | 455 |
exception THEORY of string * theory list; |
456 |
||
457 |
||
387
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
458 |
val sign_of = #sign o rep_theory; |
0 | 459 |
val syn_of = #syn o Sign.rep_sg o sign_of; |
460 |
||
387
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
461 |
(*stamps associated with a theory*) |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
462 |
val stamps_of_thy = #stamps o Sign.rep_sg o sign_of; |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
463 |
|
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
464 |
(*return the immediate ancestors*) |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
465 |
val parents_of = #parents o rep_theory; |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
466 |
|
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
467 |
|
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
468 |
(*compare theories*) |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
469 |
val subthy = Sign.subsig o pairself sign_of; |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
470 |
val eq_thy = Sign.eq_sg o pairself sign_of; |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
471 |
|
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
472 |
|
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
473 |
(*look up the named axiom in the theory*) |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
474 |
fun get_axiom theory name = |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
475 |
let |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
476 |
fun get_ax [] = raise Match |
399
86cc2b98f9e0
added class_triv: theory -> class -> thm (for axclasses);
wenzelm
parents:
387
diff
changeset
|
477 |
| get_ax (Theory {sign, new_axioms, parents} :: thys) = |
86cc2b98f9e0
added class_triv: theory -> class -> thm (for axclasses);
wenzelm
parents:
387
diff
changeset
|
478 |
(case Symtab.lookup (new_axioms, name) of |
1238 | 479 |
Some t => fix_shyps [] [] |
480 |
(Thm {sign = sign, maxidx = maxidx_of_term t, |
|
481 |
shyps = [], hyps = [], prop = t}) |
|
387
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
482 |
| None => get_ax parents handle Match => get_ax thys); |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
483 |
in |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
484 |
get_ax [theory] handle Match |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
485 |
=> raise THEORY ("get_axiom: no axiom " ^ quote name, [theory]) |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
486 |
end; |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
487 |
|
776
df8f91c0e57c
improved axioms_of: returns thms as the manual says;
wenzelm
parents:
721
diff
changeset
|
488 |
(*return additional axioms of this theory node*) |
df8f91c0e57c
improved axioms_of: returns thms as the manual says;
wenzelm
parents:
721
diff
changeset
|
489 |
fun axioms_of thy = |
df8f91c0e57c
improved axioms_of: returns thms as the manual says;
wenzelm
parents:
721
diff
changeset
|
490 |
map (fn (s, _) => (s, get_axiom thy s)) |
df8f91c0e57c
improved axioms_of: returns thms as the manual says;
wenzelm
parents:
721
diff
changeset
|
491 |
(Symtab.dest (#new_axioms (rep_theory thy))); |
df8f91c0e57c
improved axioms_of: returns thms as the manual says;
wenzelm
parents:
721
diff
changeset
|
492 |
|
387
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
493 |
|
922
196ca0973a6d
added CPure (curried functions) and ProtoPure (ancestor of Pure and CPure)
clasohm
parents:
900
diff
changeset
|
494 |
(* the Pure theories *) |
196ca0973a6d
added CPure (curried functions) and ProtoPure (ancestor of Pure and CPure)
clasohm
parents:
900
diff
changeset
|
495 |
|
196ca0973a6d
added CPure (curried functions) and ProtoPure (ancestor of Pure and CPure)
clasohm
parents:
900
diff
changeset
|
496 |
val proto_pure_thy = |
196ca0973a6d
added CPure (curried functions) and ProtoPure (ancestor of Pure and CPure)
clasohm
parents:
900
diff
changeset
|
497 |
Theory {sign = Sign.proto_pure, new_axioms = Symtab.null, parents = []}; |
387
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
498 |
|
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
499 |
val pure_thy = |
399
86cc2b98f9e0
added class_triv: theory -> class -> thm (for axclasses);
wenzelm
parents:
387
diff
changeset
|
500 |
Theory {sign = Sign.pure, new_axioms = Symtab.null, parents = []}; |
387
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
501 |
|
922
196ca0973a6d
added CPure (curried functions) and ProtoPure (ancestor of Pure and CPure)
clasohm
parents:
900
diff
changeset
|
502 |
val cpure_thy = |
196ca0973a6d
added CPure (curried functions) and ProtoPure (ancestor of Pure and CPure)
clasohm
parents:
900
diff
changeset
|
503 |
Theory {sign = Sign.cpure, new_axioms = Symtab.null, parents = []}; |
196ca0973a6d
added CPure (curried functions) and ProtoPure (ancestor of Pure and CPure)
clasohm
parents:
900
diff
changeset
|
504 |
|
0 | 505 |
|
387
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
506 |
|
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
507 |
(** extend theory **) |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
508 |
|
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
509 |
fun err_dup_axms names = |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
510 |
error ("Duplicate axiom name(s) " ^ commas_quote names); |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
511 |
|
399
86cc2b98f9e0
added class_triv: theory -> class -> thm (for axclasses);
wenzelm
parents:
387
diff
changeset
|
512 |
fun ext_thy (thy as Theory {sign, new_axioms, parents}) sign1 new_axms = |
387
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
513 |
let |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
514 |
val draft = Sign.is_draft sign; |
399
86cc2b98f9e0
added class_triv: theory -> class -> thm (for axclasses);
wenzelm
parents:
387
diff
changeset
|
515 |
val new_axioms1 = |
86cc2b98f9e0
added class_triv: theory -> class -> thm (for axclasses);
wenzelm
parents:
387
diff
changeset
|
516 |
Symtab.extend_new (if draft then new_axioms else Symtab.null, new_axms) |
387
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
517 |
handle Symtab.DUPS names => err_dup_axms names; |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
518 |
val parents1 = if draft then parents else [thy]; |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
519 |
in |
399
86cc2b98f9e0
added class_triv: theory -> class -> thm (for axclasses);
wenzelm
parents:
387
diff
changeset
|
520 |
Theory {sign = sign1, new_axioms = new_axioms1, parents = parents1} |
387
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
521 |
end; |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
522 |
|
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
523 |
|
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
524 |
(* extend signature of a theory *) |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
525 |
|
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
526 |
fun ext_sg extfun decls (thy as Theory {sign, ...}) = |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
527 |
ext_thy thy (extfun decls sign) []; |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
528 |
|
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
529 |
val add_classes = ext_sg Sign.add_classes; |
421 | 530 |
val add_classrel = ext_sg Sign.add_classrel; |
387
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
531 |
val add_defsort = ext_sg Sign.add_defsort; |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
532 |
val add_types = ext_sg Sign.add_types; |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
533 |
val add_tyabbrs = ext_sg Sign.add_tyabbrs; |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
534 |
val add_tyabbrs_i = ext_sg Sign.add_tyabbrs_i; |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
535 |
val add_arities = ext_sg Sign.add_arities; |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
536 |
val add_consts = ext_sg Sign.add_consts; |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
537 |
val add_consts_i = ext_sg Sign.add_consts_i; |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
538 |
val add_syntax = ext_sg Sign.add_syntax; |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
539 |
val add_syntax_i = ext_sg Sign.add_syntax_i; |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
540 |
val add_trfuns = ext_sg Sign.add_trfuns; |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
541 |
val add_trrules = ext_sg Sign.add_trrules; |
1160 | 542 |
val add_trrules_i = ext_sg Sign.add_trrules_i; |
387
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
543 |
val add_thyname = ext_sg Sign.add_name; |
0 | 544 |
|
545 |
||
387
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
546 |
(* prepare axioms *) |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
547 |
|
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
548 |
fun err_in_axm name = |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
549 |
error ("The error(s) above occurred in axiom " ^ quote name); |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
550 |
|
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
551 |
fun no_vars tm = |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
552 |
if null (term_vars tm) andalso null (term_tvars tm) then tm |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
553 |
else error "Illegal schematic variable(s) in term"; |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
554 |
|
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
555 |
fun cert_axm sg (name, raw_tm) = |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
556 |
let |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
557 |
val Cterm {t, T, ...} = cterm_of sg raw_tm |
1394
a1d2735f5ade
New function read_cterms is a combination of read_def_cterm and
paulson
parents:
1258
diff
changeset
|
558 |
handle TYPE arg => error (Sign.exn_type_msg sg arg) |
1460 | 559 |
| TERM (msg, _) => error msg; |
387
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
560 |
in |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
561 |
assert (T = propT) "Term not of type prop"; |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
562 |
(name, no_vars t) |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
563 |
end |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
564 |
handle ERROR => err_in_axm name; |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
565 |
|
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
566 |
fun read_axm sg (name, str) = |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
567 |
(name, no_vars (term_of (read_cterm sg (str, propT)))) |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
568 |
handle ERROR => err_in_axm name; |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
569 |
|
564 | 570 |
fun inferT_axm sg (name, pre_tm) = |
959 | 571 |
let val t = #2(Sign.infer_types sg (K None) (K None) [] true |
949
83c588d6fee9
Changed treatment of during type inference internally generated type
nipkow
parents:
922
diff
changeset
|
572 |
([pre_tm], propT)) |
83c588d6fee9
Changed treatment of during type inference internally generated type
nipkow
parents:
922
diff
changeset
|
573 |
in (name, no_vars t) end |
83c588d6fee9
Changed treatment of during type inference internally generated type
nipkow
parents:
922
diff
changeset
|
574 |
handle ERROR => err_in_axm name; |
564 | 575 |
|
387
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
576 |
|
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
577 |
(* extend axioms of a theory *) |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
578 |
|
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
579 |
fun ext_axms prep_axm axms (thy as Theory {sign, ...}) = |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
580 |
let |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
581 |
val sign1 = Sign.make_draft sign; |
1416 | 582 |
val axioms = map (apsnd (Term.compress_term o Logic.varify) o |
1460 | 583 |
prep_axm sign) |
584 |
axms; |
|
387
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
585 |
in |
399
86cc2b98f9e0
added class_triv: theory -> class -> thm (for axclasses);
wenzelm
parents:
387
diff
changeset
|
586 |
ext_thy thy sign1 axioms |
387
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
587 |
end; |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
588 |
|
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
589 |
val add_axioms = ext_axms read_axm; |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
590 |
val add_axioms_i = ext_axms cert_axm; |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
591 |
|
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
592 |
|
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
593 |
|
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
594 |
(** merge theories **) |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
595 |
|
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
596 |
fun merge_thy_list mk_draft thys = |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
597 |
let |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
598 |
fun is_union thy = forall (fn t => subthy (t, thy)) thys; |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
599 |
val is_draft = Sign.is_draft o sign_of; |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
600 |
|
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
601 |
fun add_sign (sg, Theory {sign, ...}) = |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
602 |
Sign.merge (sg, sign) handle TERM (msg, _) => error msg; |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
603 |
in |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
604 |
(case (find_first is_union thys, exists is_draft thys) of |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
605 |
(Some thy, _) => thy |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
606 |
| (None, true) => raise THEORY ("Illegal merge of draft theories", thys) |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
607 |
| (None, false) => Theory { |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
608 |
sign = |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
609 |
(if mk_draft then Sign.make_draft else I) |
922
196ca0973a6d
added CPure (curried functions) and ProtoPure (ancestor of Pure and CPure)
clasohm
parents:
900
diff
changeset
|
610 |
(foldl add_sign (Sign.proto_pure, thys)), |
399
86cc2b98f9e0
added class_triv: theory -> class -> thm (for axclasses);
wenzelm
parents:
387
diff
changeset
|
611 |
new_axioms = Symtab.null, |
387
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
612 |
parents = thys}) |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
613 |
end; |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
614 |
|
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
615 |
fun merge_theories (thy1, thy2) = merge_thy_list false [thy1, thy2]; |
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
616 |
|
0 | 617 |
|
618 |
||
1220 | 619 |
(*** Meta rules ***) |
620 |
||
621 |
(** 'primitive' rules **) |
|
622 |
||
623 |
(*discharge all assumptions t from ts*) |
|
0 | 624 |
val disch = gen_rem (op aconv); |
625 |
||
1220 | 626 |
(*The assumption rule A|-A in a theory*) |
250 | 627 |
fun assume ct : thm = |
229
4002c4cd450c
Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents:
225
diff
changeset
|
628 |
let val {sign, t=prop, T, maxidx} = rep_cterm ct |
250 | 629 |
in if T<>propT then |
630 |
raise THM("assume: assumptions must have type prop", 0, []) |
|
0 | 631 |
else if maxidx <> ~1 then |
250 | 632 |
raise THM("assume: assumptions may not contain scheme variables", |
633 |
maxidx, []) |
|
1238 | 634 |
else fix_shyps [] [] |
635 |
(Thm{sign = sign, maxidx = ~1, shyps = [], hyps = [prop], prop = prop}) |
|
0 | 636 |
end; |
637 |
||
1220 | 638 |
(*Implication introduction |
639 |
A |- B |
|
640 |
------- |
|
641 |
A ==> B |
|
642 |
*) |
|
1238 | 643 |
fun implies_intr cA (thB as Thm{sign,maxidx,hyps,prop,...}) : thm = |
229
4002c4cd450c
Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents:
225
diff
changeset
|
644 |
let val {sign=signA, t=A, T, maxidx=maxidxA} = rep_cterm cA |
0 | 645 |
in if T<>propT then |
250 | 646 |
raise THM("implies_intr: assumptions must have type prop", 0, [thB]) |
1238 | 647 |
else fix_shyps [thB] [] |
648 |
(Thm{sign= Sign.merge (sign,signA), maxidx= max[maxidxA, maxidx], |
|
649 |
shyps= [], hyps= disch(hyps,A), prop= implies$A$prop}) |
|
0 | 650 |
handle TERM _ => |
651 |
raise THM("implies_intr: incompatible signatures", 0, [thB]) |
|
652 |
end; |
|
653 |
||
1220 | 654 |
(*Implication elimination |
655 |
A ==> B A |
|
656 |
------------ |
|
657 |
B |
|
658 |
*) |
|
0 | 659 |
fun implies_elim thAB thA : thm = |
660 |
let val Thm{maxidx=maxA, hyps=hypsA, prop=propA,...} = thA |
|
250 | 661 |
and Thm{sign, maxidx, hyps, prop,...} = thAB; |
662 |
fun err(a) = raise THM("implies_elim: "^a, 0, [thAB,thA]) |
|
0 | 663 |
in case prop of |
250 | 664 |
imp$A$B => |
665 |
if imp=implies andalso A aconv propA |
|
1220 | 666 |
then fix_shyps [thAB, thA] [] |
667 |
(Thm{sign= merge_thm_sgs(thAB,thA), |
|
250 | 668 |
maxidx= max[maxA,maxidx], |
1220 | 669 |
shyps= [], |
250 | 670 |
hyps= hypsA union hyps, (*dups suppressed*) |
1220 | 671 |
prop= B}) |
250 | 672 |
else err("major premise") |
673 |
| _ => err("major premise") |
|
0 | 674 |
end; |
250 | 675 |
|
1220 | 676 |
(*Forall introduction. The Free or Var x must not be free in the hypotheses. |
677 |
A |
|
678 |
----- |
|
679 |
!!x.A |
|
680 |
*) |
|
1238 | 681 |
fun forall_intr cx (th as Thm{sign,maxidx,hyps,prop,...}) = |
229
4002c4cd450c
Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents:
225
diff
changeset
|
682 |
let val x = term_of cx; |
1238 | 683 |
fun result(a,T) = fix_shyps [th] [] |
684 |
(Thm{sign= sign, maxidx= maxidx, shyps= [], hyps= hyps, |
|
685 |
prop= all(T) $ Abs(a, T, abstract_over (x,prop))}) |
|
0 | 686 |
in case x of |
250 | 687 |
Free(a,T) => |
688 |
if exists (apl(x, Logic.occs)) hyps |
|
689 |
then raise THM("forall_intr: variable free in assumptions", 0, [th]) |
|
690 |
else result(a,T) |
|
0 | 691 |
| Var((a,_),T) => result(a,T) |
692 |
| _ => raise THM("forall_intr: not a variable", 0, [th]) |
|
693 |
end; |
|
694 |
||
1220 | 695 |
(*Forall elimination |
696 |
!!x.A |
|
697 |
------ |
|
698 |
A[t/x] |
|
699 |
*) |
|
700 |
fun forall_elim ct (th as Thm{sign,maxidx,hyps,prop,...}) : thm = |
|
229
4002c4cd450c
Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents:
225
diff
changeset
|
701 |
let val {sign=signt, t, T, maxidx=maxt} = rep_cterm ct |
0 | 702 |
in case prop of |
250 | 703 |
Const("all",Type("fun",[Type("fun",[qary,_]),_])) $ A => |
704 |
if T<>qary then |
|
705 |
raise THM("forall_elim: type mismatch", 0, [th]) |
|
1220 | 706 |
else fix_shyps [th] [] |
707 |
(Thm{sign= Sign.merge(sign,signt), |
|
250 | 708 |
maxidx= max[maxidx, maxt], |
1220 | 709 |
shyps= [], hyps= hyps, prop= betapply(A,t)}) |
250 | 710 |
| _ => raise THM("forall_elim: not quantified", 0, [th]) |
0 | 711 |
end |
712 |
handle TERM _ => |
|
250 | 713 |
raise THM("forall_elim: incompatible signatures", 0, [th]); |
0 | 714 |
|
715 |
||
1220 | 716 |
(* Equality *) |
0 | 717 |
|
1220 | 718 |
(* Definition of the relation =?= *) |
1238 | 719 |
val flexpair_def = fix_shyps [] [] |
720 |
(Thm{sign= Sign.proto_pure, shyps= [], hyps= [], maxidx= 0, |
|
721 |
prop= term_of (read_cterm Sign.proto_pure |
|
722 |
("(?t =?= ?u) == (?t == ?u::?'a::{})", propT))}); |
|
0 | 723 |
|
724 |
(*The reflexivity rule: maps t to the theorem t==t *) |
|
250 | 725 |
fun reflexive ct = |
229
4002c4cd450c
Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents:
225
diff
changeset
|
726 |
let val {sign, t, T, maxidx} = rep_cterm ct |
1238 | 727 |
in fix_shyps [] [] |
728 |
(Thm{sign= sign, shyps= [], hyps= [], maxidx= maxidx, |
|
729 |
prop= Logic.mk_equals(t,t)}) |
|
0 | 730 |
end; |
731 |
||
732 |
(*The symmetry rule |
|
1220 | 733 |
t==u |
734 |
---- |
|
735 |
u==t |
|
736 |
*) |
|
737 |
fun symmetric (th as Thm{sign,shyps,hyps,prop,maxidx}) = |
|
0 | 738 |
case prop of |
739 |
(eq as Const("==",_)) $ t $ u => |
|
1238 | 740 |
(*no fix_shyps*) |
741 |
Thm{sign=sign, shyps=shyps, hyps=hyps, maxidx=maxidx, prop= eq$u$t} |
|
0 | 742 |
| _ => raise THM("symmetric", 0, [th]); |
743 |
||
744 |
(*The transitive rule |
|
1220 | 745 |
t1==u u==t2 |
746 |
-------------- |
|
747 |
t1==t2 |
|
748 |
*) |
|
0 | 749 |
fun transitive th1 th2 = |
750 |
let val Thm{maxidx=max1, hyps=hyps1, prop=prop1,...} = th1 |
|
751 |
and Thm{maxidx=max2, hyps=hyps2, prop=prop2,...} = th2; |
|
752 |
fun err(msg) = raise THM("transitive: "^msg, 0, [th1,th2]) |
|
753 |
in case (prop1,prop2) of |
|
754 |
((eq as Const("==",_)) $ t1 $ u, Const("==",_) $ u' $ t2) => |
|
250 | 755 |
if not (u aconv u') then err"middle term" else |
1220 | 756 |
fix_shyps [th1, th2] [] |
757 |
(Thm{sign= merge_thm_sgs(th1,th2), shyps= [], |
|
758 |
hyps= hyps1 union hyps2, |
|
759 |
maxidx= max[max1,max2], prop= eq$t1$t2}) |
|
0 | 760 |
| _ => err"premises" |
761 |
end; |
|
762 |
||
1160 | 763 |
(*Beta-conversion: maps (%x.t)(u) to the theorem (%x.t)(u) == t[u/x] *) |
250 | 764 |
fun beta_conversion ct = |
229
4002c4cd450c
Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents:
225
diff
changeset
|
765 |
let val {sign, t, T, maxidx} = rep_cterm ct |
0 | 766 |
in case t of |
1238 | 767 |
Abs(_,_,bodt) $ u => fix_shyps [] [] |
768 |
(Thm{sign= sign, shyps= [], hyps= [], |
|
250 | 769 |
maxidx= maxidx_of_term t, |
1238 | 770 |
prop= Logic.mk_equals(t, subst_bounds([u],bodt))}) |
250 | 771 |
| _ => raise THM("beta_conversion: not a redex", 0, []) |
0 | 772 |
end; |
773 |
||
774 |
(*The extensionality rule (proviso: x not free in f, g, or hypotheses) |
|
1220 | 775 |
f(x) == g(x) |
776 |
------------ |
|
777 |
f == g |
|
778 |
*) |
|
779 |
fun extensional (th as Thm{sign,maxidx,shyps,hyps,prop}) = |
|
0 | 780 |
case prop of |
781 |
(Const("==",_)) $ (f$x) $ (g$y) => |
|
250 | 782 |
let fun err(msg) = raise THM("extensional: "^msg, 0, [th]) |
0 | 783 |
in (if x<>y then err"different variables" else |
784 |
case y of |
|
250 | 785 |
Free _ => |
786 |
if exists (apl(y, Logic.occs)) (f::g::hyps) |
|
787 |
then err"variable free in hyps or functions" else () |
|
788 |
| Var _ => |
|
789 |
if Logic.occs(y,f) orelse Logic.occs(y,g) |
|
790 |
then err"variable free in functions" else () |
|
791 |
| _ => err"not a variable"); |
|
1238 | 792 |
(*no fix_shyps*) |
1220 | 793 |
Thm{sign=sign, shyps=shyps, hyps=hyps, maxidx=maxidx, |
250 | 794 |
prop= Logic.mk_equals(f,g)} |
0 | 795 |
end |
796 |
| _ => raise THM("extensional: premise", 0, [th]); |
|
797 |
||
798 |
(*The abstraction rule. The Free or Var x must not be free in the hypotheses. |
|
799 |
The bound variable will be named "a" (since x will be something like x320) |
|
1220 | 800 |
t == u |
801 |
------------ |
|
802 |
%x.t == %x.u |
|
803 |
*) |
|
1238 | 804 |
fun abstract_rule a cx (th as Thm{sign,maxidx,hyps,prop,...}) = |
229
4002c4cd450c
Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents:
225
diff
changeset
|
805 |
let val x = term_of cx; |
250 | 806 |
val (t,u) = Logic.dest_equals prop |
807 |
handle TERM _ => |
|
808 |
raise THM("abstract_rule: premise not an equality", 0, [th]) |
|
1238 | 809 |
fun result T = fix_shyps [th] [] |
810 |
(Thm{sign= sign, maxidx= maxidx, shyps= [], hyps= hyps, |
|
250 | 811 |
prop= Logic.mk_equals(Abs(a, T, abstract_over (x,t)), |
1238 | 812 |
Abs(a, T, abstract_over (x,u)))}) |
0 | 813 |
in case x of |
250 | 814 |
Free(_,T) => |
815 |
if exists (apl(x, Logic.occs)) hyps |
|
816 |
then raise THM("abstract_rule: variable free in assumptions", 0, [th]) |
|
817 |
else result T |
|
0 | 818 |
| Var(_,T) => result T |
819 |
| _ => raise THM("abstract_rule: not a variable", 0, [th]) |
|
820 |
end; |
|
821 |
||
822 |
(*The combination rule |
|
1220 | 823 |
f==g t==u |
824 |
------------ |
|
825 |
f(t)==g(u) |
|
826 |
*) |
|
0 | 827 |
fun combination th1 th2 = |
1220 | 828 |
let val Thm{maxidx=max1, shyps=shyps1, hyps=hyps1, prop=prop1,...} = th1 |
829 |
and Thm{maxidx=max2, shyps=shyps2, hyps=hyps2, prop=prop2,...} = th2 |
|
0 | 830 |
in case (prop1,prop2) of |
831 |
(Const("==",_) $ f $ g, Const("==",_) $ t $ u) => |
|
1238 | 832 |
(*no fix_shyps*) |
1220 | 833 |
Thm{sign= merge_thm_sgs(th1,th2), shyps= shyps1 union shyps2, |
834 |
hyps= hyps1 union hyps2, |
|
250 | 835 |
maxidx= max[max1,max2], prop= Logic.mk_equals(f$t, g$u)} |
0 | 836 |
| _ => raise THM("combination: premises", 0, [th1,th2]) |
837 |
end; |
|
838 |
||
839 |
||
840 |
(*The equal propositions rule |
|
1220 | 841 |
A==B A |
842 |
--------- |
|
843 |
B |
|
844 |
*) |
|
0 | 845 |
fun equal_elim th1 th2 = |
846 |
let val Thm{maxidx=max1, hyps=hyps1, prop=prop1,...} = th1 |
|
847 |
and Thm{maxidx=max2, hyps=hyps2, prop=prop2,...} = th2; |
|
848 |
fun err(msg) = raise THM("equal_elim: "^msg, 0, [th1,th2]) |
|
849 |
in case prop1 of |
|
850 |
Const("==",_) $ A $ B => |
|
250 | 851 |
if not (prop2 aconv A) then err"not equal" else |
1220 | 852 |
fix_shyps [th1, th2] [] |
853 |
(Thm{sign= merge_thm_sgs(th1,th2), shyps= [], |
|
854 |
hyps= hyps1 union hyps2, |
|
855 |
maxidx= max[max1,max2], prop= B}) |
|
0 | 856 |
| _ => err"major premise" |
857 |
end; |
|
858 |
||
859 |
||
860 |
(* Equality introduction |
|
1220 | 861 |
A==>B B==>A |
862 |
-------------- |
|
863 |
A==B |
|
864 |
*) |
|
0 | 865 |
fun equal_intr th1 th2 = |
1220 | 866 |
let val Thm{maxidx=max1, shyps=shyps1, hyps=hyps1, prop=prop1,...} = th1 |
867 |
and Thm{maxidx=max2, shyps=shyps2, hyps=hyps2, prop=prop2,...} = th2; |
|
0 | 868 |
fun err(msg) = raise THM("equal_intr: "^msg, 0, [th1,th2]) |
869 |
in case (prop1,prop2) of |
|
870 |
(Const("==>",_) $ A $ B, Const("==>",_) $ B' $ A') => |
|
250 | 871 |
if A aconv A' andalso B aconv B' |
1238 | 872 |
then |
873 |
(*no fix_shyps*) |
|
874 |
Thm{sign= merge_thm_sgs(th1,th2), shyps= shyps1 union shyps2, |
|
875 |
hyps= hyps1 union hyps2, |
|
876 |
maxidx= max[max1,max2], prop= Logic.mk_equals(A,B)} |
|
250 | 877 |
else err"not equal" |
0 | 878 |
| _ => err"premises" |
879 |
end; |
|
880 |
||
1220 | 881 |
|
882 |
||
0 | 883 |
(**** Derived rules ****) |
884 |
||
885 |
(*Discharge all hypotheses (need not verify cterms) |
|
886 |
Repeated hypotheses are discharged only once; fold cannot do this*) |
|
1220 | 887 |
fun implies_intr_hyps (Thm{sign, maxidx, shyps, hyps=A::As, prop}) = |
1238 | 888 |
implies_intr_hyps (*no fix_shyps*) |
1220 | 889 |
(Thm{sign=sign, maxidx=maxidx, shyps=shyps, |
250 | 890 |
hyps= disch(As,A), prop= implies$A$prop}) |
0 | 891 |
| implies_intr_hyps th = th; |
892 |
||
893 |
(*Smash" unifies the list of term pairs leaving no flex-flex pairs. |
|
250 | 894 |
Instantiates the theorem and deletes trivial tpairs. |
0 | 895 |
Resulting sequence may contain multiple elements if the tpairs are |
896 |
not all flex-flex. *) |
|
1220 | 897 |
fun flexflex_rule (th as Thm{sign,maxidx,hyps,prop,...}) = |
250 | 898 |
let fun newthm env = |
899 |
let val (tpairs,horn) = |
|
900 |
Logic.strip_flexpairs (Envir.norm_term env prop) |
|
901 |
(*Remove trivial tpairs, of the form t=t*) |
|
902 |
val distpairs = filter (not o op aconv) tpairs |
|
903 |
val newprop = Logic.list_flexpairs(distpairs, horn) |
|
1220 | 904 |
in fix_shyps [th] (env_codT env) |
905 |
(Thm{sign= sign, shyps= [], hyps= hyps, |
|
906 |
maxidx= maxidx_of_term newprop, prop= newprop}) |
|
250 | 907 |
end; |
0 | 908 |
val (tpairs,_) = Logic.strip_flexpairs prop |
909 |
in Sequence.maps newthm |
|
250 | 910 |
(Unify.smash_unifiers(sign, Envir.empty maxidx, tpairs)) |
0 | 911 |
end; |
912 |
||
913 |
(*Instantiation of Vars |
|
1220 | 914 |
A |
915 |
------------------- |
|
916 |
A[t1/v1,....,tn/vn] |
|
917 |
*) |
|
0 | 918 |
|
919 |
(*Check that all the terms are Vars and are distinct*) |
|
920 |
fun instl_ok ts = forall is_Var ts andalso null(findrep ts); |
|
921 |
||
922 |
(*For instantiate: process pair of cterms, merge theories*) |
|
923 |
fun add_ctpair ((ct,cu), (sign,tpairs)) = |
|
229
4002c4cd450c
Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents:
225
diff
changeset
|
924 |
let val {sign=signt, t=t, T= T, ...} = rep_cterm ct |
4002c4cd450c
Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents:
225
diff
changeset
|
925 |
and {sign=signu, t=u, T= U, ...} = rep_cterm cu |
0 | 926 |
in if T=U then (Sign.merge(sign, Sign.merge(signt, signu)), (t,u)::tpairs) |
927 |
else raise TYPE("add_ctpair", [T,U], [t,u]) |
|
928 |
end; |
|
929 |
||
930 |
fun add_ctyp ((v,ctyp), (sign',vTs)) = |
|
229
4002c4cd450c
Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents:
225
diff
changeset
|
931 |
let val {T,sign} = rep_ctyp ctyp |
0 | 932 |
in (Sign.merge(sign,sign'), (v,T)::vTs) end; |
933 |
||
934 |
(*Left-to-right replacements: ctpairs = [...,(vi,ti),...]. |
|
935 |
Instantiates distinct Vars by terms of same type. |
|
936 |
Normalizes the new theorem! *) |
|
1220 | 937 |
fun instantiate (vcTs,ctpairs) (th as Thm{sign,maxidx,hyps,prop,...}) = |
0 | 938 |
let val (newsign,tpairs) = foldr add_ctpair (ctpairs, (sign,[])); |
939 |
val (newsign,vTs) = foldr add_ctyp (vcTs, (newsign,[])); |
|
250 | 940 |
val newprop = |
941 |
Envir.norm_term (Envir.empty 0) |
|
942 |
(subst_atomic tpairs |
|
943 |
(Type.inst_term_tvars(#tsig(Sign.rep_sg newsign),vTs) prop)) |
|
1220 | 944 |
val newth = |
945 |
fix_shyps [th] (map snd vTs) |
|
946 |
(Thm{sign= newsign, shyps= [], hyps= hyps, |
|
947 |
maxidx= maxidx_of_term newprop, prop= newprop}) |
|
250 | 948 |
in if not(instl_ok(map #1 tpairs)) |
193 | 949 |
then raise THM("instantiate: variables not distinct", 0, [th]) |
950 |
else if not(null(findrep(map #1 vTs))) |
|
951 |
then raise THM("instantiate: type variables not distinct", 0, [th]) |
|
952 |
else (*Check types of Vars for agreement*) |
|
953 |
case findrep (map (#1 o dest_Var) (term_vars newprop)) of |
|
250 | 954 |
ix::_ => raise THM("instantiate: conflicting types for variable " ^ |
955 |
Syntax.string_of_vname ix ^ "\n", 0, [newth]) |
|
956 |
| [] => |
|
957 |
case findrep (map #1 (term_tvars newprop)) of |
|
958 |
ix::_ => raise THM |
|
959 |
("instantiate: conflicting sorts for type variable " ^ |
|
960 |
Syntax.string_of_vname ix ^ "\n", 0, [newth]) |
|
193 | 961 |
| [] => newth |
0 | 962 |
end |
250 | 963 |
handle TERM _ => |
0 | 964 |
raise THM("instantiate: incompatible signatures",0,[th]) |
193 | 965 |
| TYPE _ => raise THM("instantiate: type conflict", 0, [th]); |
0 | 966 |
|
967 |
(*The trivial implication A==>A, justified by assume and forall rules. |
|
968 |
A can contain Vars, not so for assume! *) |
|
250 | 969 |
fun trivial ct : thm = |
229
4002c4cd450c
Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents:
225
diff
changeset
|
970 |
let val {sign, t=A, T, maxidx} = rep_cterm ct |
250 | 971 |
in if T<>propT then |
972 |
raise THM("trivial: the term must have type prop", 0, []) |
|
1238 | 973 |
else fix_shyps [] [] |
974 |
(Thm{sign= sign, maxidx= maxidx, shyps= [], hyps= [], |
|
975 |
prop= implies$A$A}) |
|
0 | 976 |
end; |
977 |
||
1160 | 978 |
(*Axiom-scheme reflecting signature contents: "OFCLASS(?'a::c, c_class)" -- |
1220 | 979 |
essentially just an instance of A==>A.*) |
399
86cc2b98f9e0
added class_triv: theory -> class -> thm (for axclasses);
wenzelm
parents:
387
diff
changeset
|
980 |
fun class_triv thy c = |
86cc2b98f9e0
added class_triv: theory -> class -> thm (for axclasses);
wenzelm
parents:
387
diff
changeset
|
981 |
let |
86cc2b98f9e0
added class_triv: theory -> class -> thm (for axclasses);
wenzelm
parents:
387
diff
changeset
|
982 |
val sign = sign_of thy; |
86cc2b98f9e0
added class_triv: theory -> class -> thm (for axclasses);
wenzelm
parents:
387
diff
changeset
|
983 |
val Cterm {t, maxidx, ...} = |
86cc2b98f9e0
added class_triv: theory -> class -> thm (for axclasses);
wenzelm
parents:
387
diff
changeset
|
984 |
cterm_of sign (Logic.mk_inclass (TVar (("'a", 0), [c]), c)) |
86cc2b98f9e0
added class_triv: theory -> class -> thm (for axclasses);
wenzelm
parents:
387
diff
changeset
|
985 |
handle TERM (msg, _) => raise THM ("class_triv: " ^ msg, 0, []); |
86cc2b98f9e0
added class_triv: theory -> class -> thm (for axclasses);
wenzelm
parents:
387
diff
changeset
|
986 |
in |
1238 | 987 |
fix_shyps [] [] |
988 |
(Thm {sign = sign, maxidx = maxidx, shyps = [], hyps = [], prop = t}) |
|
399
86cc2b98f9e0
added class_triv: theory -> class -> thm (for axclasses);
wenzelm
parents:
387
diff
changeset
|
989 |
end; |
86cc2b98f9e0
added class_triv: theory -> class -> thm (for axclasses);
wenzelm
parents:
387
diff
changeset
|
990 |
|
86cc2b98f9e0
added class_triv: theory -> class -> thm (for axclasses);
wenzelm
parents:
387
diff
changeset
|
991 |
|
0 | 992 |
(* Replace all TFrees not in the hyps by new TVars *) |
1220 | 993 |
fun varifyT(Thm{sign,maxidx,shyps,hyps,prop}) = |
0 | 994 |
let val tfrees = foldr add_term_tfree_names (hyps,[]) |
1238 | 995 |
in (*no fix_shyps*) |
996 |
Thm{sign=sign, maxidx=max[0,maxidx], shyps=shyps, hyps=hyps, |
|
997 |
prop= Type.varify(prop,tfrees)} |
|
0 | 998 |
end; |
999 |
||
1000 |
(* Replace all TVars by new TFrees *) |
|
1220 | 1001 |
fun freezeT(Thm{sign,maxidx,shyps,hyps,prop}) = |
949
83c588d6fee9
Changed treatment of during type inference internally generated type
nipkow
parents:
922
diff
changeset
|
1002 |
let val prop' = Type.freeze prop |
1238 | 1003 |
in (*no fix_shyps*) |
1004 |
Thm{sign=sign, maxidx=maxidx_of_term prop', shyps=shyps, hyps=hyps, |
|
1005 |
prop=prop'} |
|
1220 | 1006 |
end; |
0 | 1007 |
|
1008 |
||
1009 |
(*** Inference rules for tactics ***) |
|
1010 |
||
1011 |
(*Destruct proof state into constraints, other goals, goal(i), rest *) |
|
1012 |
fun dest_state (state as Thm{prop,...}, i) = |
|
1013 |
let val (tpairs,horn) = Logic.strip_flexpairs prop |
|
1014 |
in case Logic.strip_prems(i, [], horn) of |
|
1015 |
(B::rBs, C) => (tpairs, rev rBs, B, C) |
|
1016 |
| _ => raise THM("dest_state", i, [state]) |
|
1017 |
end |
|
1018 |
handle TERM _ => raise THM("dest_state", i, [state]); |
|
1019 |
||
309 | 1020 |
(*Increment variables and parameters of orule as required for |
0 | 1021 |
resolution with goal i of state. *) |
1022 |
fun lift_rule (state, i) orule = |
|
1238 | 1023 |
let val Thm{shyps=sshyps,prop=sprop,maxidx=smax,...} = state; |
0 | 1024 |
val (Bi::_, _) = Logic.strip_prems(i, [], Logic.skip_flexpairs sprop) |
250 | 1025 |
handle TERM _ => raise THM("lift_rule", i, [orule,state]); |
0 | 1026 |
val (lift_abs,lift_all) = Logic.lift_fns(Bi,smax+1); |
1238 | 1027 |
val (Thm{sign,maxidx,shyps,hyps,prop}) = orule |
0 | 1028 |
val (tpairs,As,B) = Logic.strip_horn prop |
1238 | 1029 |
in (*no fix_shyps*) |
1030 |
Thm{hyps=hyps, sign= merge_thm_sgs(state,orule), |
|
1031 |
shyps=sshyps union shyps, maxidx= maxidx+smax+1, |
|
250 | 1032 |
prop= Logic.rule_of(map (pairself lift_abs) tpairs, |
1238 | 1033 |
map lift_all As, lift_all B)} |
0 | 1034 |
end; |
1035 |
||
1036 |
(*Solve subgoal Bi of proof state B1...Bn/C by assumption. *) |
|
1037 |
fun assumption i state = |
|
1220 | 1038 |
let val Thm{sign,maxidx,hyps,prop,...} = state; |
0 | 1039 |
val (tpairs, Bs, Bi, C) = dest_state(state,i) |
250 | 1040 |
fun newth (env as Envir.Envir{maxidx, ...}, tpairs) = |
1220 | 1041 |
fix_shyps [state] (env_codT env) |
1042 |
(Thm{sign=sign, shyps=[], hyps=hyps, maxidx=maxidx, prop= |
|
250 | 1043 |
if Envir.is_empty env then (*avoid wasted normalizations*) |
1044 |
Logic.rule_of (tpairs, Bs, C) |
|
1045 |
else (*normalize the new rule fully*) |
|
1220 | 1046 |
Envir.norm_term env (Logic.rule_of (tpairs, Bs, C))}); |
0 | 1047 |
fun addprfs [] = Sequence.null |
1048 |
| addprfs ((t,u)::apairs) = Sequence.seqof (fn()=> Sequence.pull |
|
1049 |
(Sequence.mapp newth |
|
250 | 1050 |
(Unify.unifiers(sign,Envir.empty maxidx, (t,u)::tpairs)) |
1051 |
(addprfs apairs))) |
|
0 | 1052 |
in addprfs (Logic.assum_pairs Bi) end; |
1053 |
||
250 | 1054 |
(*Solve subgoal Bi of proof state B1...Bn/C by assumption. |
0 | 1055 |
Checks if Bi's conclusion is alpha-convertible to one of its assumptions*) |
1056 |
fun eq_assumption i state = |
|
1220 | 1057 |
let val Thm{sign,maxidx,hyps,prop,...} = state; |
0 | 1058 |
val (tpairs, Bs, Bi, C) = dest_state(state,i) |
1059 |
in if exists (op aconv) (Logic.assum_pairs Bi) |
|
1220 | 1060 |
then fix_shyps [state] [] |
1061 |
(Thm{sign=sign, shyps=[], hyps=hyps, maxidx=maxidx, |
|
1062 |
prop=Logic.rule_of(tpairs, Bs, C)}) |
|
0 | 1063 |
else raise THM("eq_assumption", 0, [state]) |
1064 |
end; |
|
1065 |
||
1066 |
||
1067 |
(** User renaming of parameters in a subgoal **) |
|
1068 |
||
1069 |
(*Calls error rather than raising an exception because it is intended |
|
1070 |
for top-level use -- exception handling would not make sense here. |
|
1071 |
The names in cs, if distinct, are used for the innermost parameters; |
|
1072 |
preceding parameters may be renamed to make all params distinct.*) |
|
1073 |
fun rename_params_rule (cs, i) state = |
|
1220 | 1074 |
let val Thm{sign,maxidx,hyps,prop,...} = state |
0 | 1075 |
val (tpairs, Bs, Bi, C) = dest_state(state,i) |
1076 |
val iparams = map #1 (Logic.strip_params Bi) |
|
1077 |
val short = length iparams - length cs |
|
250 | 1078 |
val newnames = |
1079 |
if short<0 then error"More names than abstractions!" |
|
1080 |
else variantlist(take (short,iparams), cs) @ cs |
|
0 | 1081 |
val freenames = map (#1 o dest_Free) (term_frees prop) |
1082 |
val newBi = Logic.list_rename_params (newnames, Bi) |
|
250 | 1083 |
in |
0 | 1084 |
case findrep cs of |
1085 |
c::_ => error ("Bound variables not distinct: " ^ c) |
|
1086 |
| [] => (case cs inter freenames of |
|
1087 |
a::_ => error ("Bound/Free variable clash: " ^ a) |
|
1220 | 1088 |
| [] => fix_shyps [state] [] |
1089 |
(Thm{sign=sign, shyps=[], hyps=hyps, maxidx=maxidx, prop= |
|
1090 |
Logic.rule_of(tpairs, Bs@[newBi], C)})) |
|
0 | 1091 |
end; |
1092 |
||
1093 |
(*** Preservation of bound variable names ***) |
|
1094 |
||
250 | 1095 |
(*Scan a pair of terms; while they are similar, |
0 | 1096 |
accumulate corresponding bound vars in "al"*) |
1238 | 1097 |
fun match_bvs(Abs(x,_,s),Abs(y,_,t), al) = |
1195
686e3eb613b9
match_bvs no longer puts a name in the alist if it is null ("")
lcp
parents:
1160
diff
changeset
|
1098 |
match_bvs(s, t, if x="" orelse y="" then al |
1238 | 1099 |
else (x,y)::al) |
0 | 1100 |
| match_bvs(f$s, g$t, al) = match_bvs(f,g,match_bvs(s,t,al)) |
1101 |
| match_bvs(_,_,al) = al; |
|
1102 |
||
1103 |
(* strip abstractions created by parameters *) |
|
1104 |
fun match_bvars((s,t),al) = match_bvs(strip_abs_body s, strip_abs_body t, al); |
|
1105 |
||
1106 |
||
250 | 1107 |
(* strip_apply f A(,B) strips off all assumptions/parameters from A |
0 | 1108 |
introduced by lifting over B, and applies f to remaining part of A*) |
1109 |
fun strip_apply f = |
|
1110 |
let fun strip(Const("==>",_)$ A1 $ B1, |
|
250 | 1111 |
Const("==>",_)$ _ $ B2) = implies $ A1 $ strip(B1,B2) |
1112 |
| strip((c as Const("all",_)) $ Abs(a,T,t1), |
|
1113 |
Const("all",_) $ Abs(_,_,t2)) = c$Abs(a,T,strip(t1,t2)) |
|
1114 |
| strip(A,_) = f A |
|
0 | 1115 |
in strip end; |
1116 |
||
1117 |
(*Use the alist to rename all bound variables and some unknowns in a term |
|
1118 |
dpairs = current disagreement pairs; tpairs = permanent ones (flexflex); |
|
1119 |
Preserves unknowns in tpairs and on lhs of dpairs. *) |
|
1120 |
fun rename_bvs([],_,_,_) = I |
|
1121 |
| rename_bvs(al,dpairs,tpairs,B) = |
|
250 | 1122 |
let val vars = foldr add_term_vars |
1123 |
(map fst dpairs @ map fst tpairs @ map snd tpairs, []) |
|
1124 |
(*unknowns appearing elsewhere be preserved!*) |
|
1125 |
val vids = map (#1 o #1 o dest_Var) vars; |
|
1126 |
fun rename(t as Var((x,i),T)) = |
|
1127 |
(case assoc(al,x) of |
|
1128 |
Some(y) => if x mem vids orelse y mem vids then t |
|
1129 |
else Var((y,i),T) |
|
1130 |
| None=> t) |
|
0 | 1131 |
| rename(Abs(x,T,t)) = |
250 | 1132 |
Abs(case assoc(al,x) of Some(y) => y | None => x, |
1133 |
T, rename t) |
|
0 | 1134 |
| rename(f$t) = rename f $ rename t |
1135 |
| rename(t) = t; |
|
250 | 1136 |
fun strip_ren Ai = strip_apply rename (Ai,B) |
0 | 1137 |
in strip_ren end; |
1138 |
||
1139 |
(*Function to rename bounds/unknowns in the argument, lifted over B*) |
|
1140 |
fun rename_bvars(dpairs, tpairs, B) = |
|
250 | 1141 |
rename_bvs(foldr match_bvars (dpairs,[]), dpairs, tpairs, B); |
0 | 1142 |
|
1143 |
||
1144 |
(*** RESOLUTION ***) |
|
1145 |
||
721
479832ff2d29
Pure/thm/norm_term_skip: new, for skipping normalization of the context
lcp
parents:
678
diff
changeset
|
1146 |
(** Lifting optimizations **) |
479832ff2d29
Pure/thm/norm_term_skip: new, for skipping normalization of the context
lcp
parents:
678
diff
changeset
|
1147 |
|
0 | 1148 |
(*strip off pairs of assumptions/parameters in parallel -- they are |
1149 |
identical because of lifting*) |
|
250 | 1150 |
fun strip_assums2 (Const("==>", _) $ _ $ B1, |
1151 |
Const("==>", _) $ _ $ B2) = strip_assums2 (B1,B2) |
|
0 | 1152 |
| strip_assums2 (Const("all",_)$Abs(a,T,t1), |
250 | 1153 |
Const("all",_)$Abs(_,_,t2)) = |
0 | 1154 |
let val (B1,B2) = strip_assums2 (t1,t2) |
1155 |
in (Abs(a,T,B1), Abs(a,T,B2)) end |
|
1156 |
| strip_assums2 BB = BB; |
|
1157 |
||
1158 |
||
721
479832ff2d29
Pure/thm/norm_term_skip: new, for skipping normalization of the context
lcp
parents:
678
diff
changeset
|
1159 |
(*Faster normalization: skip assumptions that were lifted over*) |
479832ff2d29
Pure/thm/norm_term_skip: new, for skipping normalization of the context
lcp
parents:
678
diff
changeset
|
1160 |
fun norm_term_skip env 0 t = Envir.norm_term env t |
479832ff2d29
Pure/thm/norm_term_skip: new, for skipping normalization of the context
lcp
parents:
678
diff
changeset
|
1161 |
| norm_term_skip env n (Const("all",_)$Abs(a,T,t)) = |
479832ff2d29
Pure/thm/norm_term_skip: new, for skipping normalization of the context
lcp
parents:
678
diff
changeset
|
1162 |
let val Envir.Envir{iTs, ...} = env |
1238 | 1163 |
val T' = typ_subst_TVars iTs T |
1164 |
(*Must instantiate types of parameters because they are flattened; |
|
721
479832ff2d29
Pure/thm/norm_term_skip: new, for skipping normalization of the context
lcp
parents:
678
diff
changeset
|
1165 |
this could be a NEW parameter*) |
479832ff2d29
Pure/thm/norm_term_skip: new, for skipping normalization of the context
lcp
parents:
678
diff
changeset
|
1166 |
in all T' $ Abs(a, T', norm_term_skip env n t) end |
479832ff2d29
Pure/thm/norm_term_skip: new, for skipping normalization of the context
lcp
parents:
678
diff
changeset
|
1167 |
| norm_term_skip env n (Const("==>", _) $ A $ B) = |
1238 | 1168 |
implies $ A $ norm_term_skip env (n-1) B |
721
479832ff2d29
Pure/thm/norm_term_skip: new, for skipping normalization of the context
lcp
parents:
678
diff
changeset
|
1169 |
| norm_term_skip env n t = error"norm_term_skip: too few assumptions??"; |
479832ff2d29
Pure/thm/norm_term_skip: new, for skipping normalization of the context
lcp
parents:
678
diff
changeset
|
1170 |
|
479832ff2d29
Pure/thm/norm_term_skip: new, for skipping normalization of the context
lcp
parents:
678
diff
changeset
|
1171 |
|
0 | 1172 |
(*Composition of object rule r=(A1...Am/B) with proof state s=(B1...Bn/C) |
250 | 1173 |
Unifies B with Bi, replacing subgoal i (1 <= i <= n) |
0 | 1174 |
If match then forbid instantiations in proof state |
1175 |
If lifted then shorten the dpair using strip_assums2. |
|
1176 |
If eres_flg then simultaneously proves A1 by assumption. |
|
250 | 1177 |
nsubgoal is the number of new subgoals (written m above). |
0 | 1178 |
Curried so that resolution calls dest_state only once. |
1179 |
*) |
|
1180 |
local open Sequence; exception Bicompose |
|
1181 |
in |
|
250 | 1182 |
fun bicompose_aux match (state, (stpairs, Bs, Bi, C), lifted) |
0 | 1183 |
(eres_flg, orule, nsubgoal) = |
1258 | 1184 |
let val Thm{maxidx=smax, shyps=sshyps, hyps=shyps, ...} = state |
1185 |
and Thm{maxidx=rmax, shyps=rshyps, hyps=rhyps, prop=rprop,...} = orule |
|
1238 | 1186 |
(*How many hyps to skip over during normalization*) |
1187 |
and nlift = Logic.count_prems(strip_all_body Bi, |
|
1188 |
if eres_flg then ~1 else 0) |
|
387
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
wenzelm
parents:
309
diff
changeset
|
1189 |
val sign = merge_thm_sgs(state,orule); |
0 | 1190 |
(** Add new theorem with prop = '[| Bs; As |] ==> C' to thq **) |
250 | 1191 |
fun addth As ((env as Envir.Envir {maxidx, ...}, tpairs), thq) = |
1192 |
let val normt = Envir.norm_term env; |
|
1193 |
(*perform minimal copying here by examining env*) |
|
1194 |
val normp = |
|
1195 |
if Envir.is_empty env then (tpairs, Bs @ As, C) |
|
1196 |
else |
|
1197 |
let val ntps = map (pairself normt) tpairs |
|
1238 | 1198 |
in if the (Envir.minidx env) > smax then |
1199 |
(*no assignments in state; normalize the rule only*) |
|
1200 |
if lifted |
|
1201 |
then (ntps, Bs @ map (norm_term_skip env nlift) As, C) |
|
1202 |
else (ntps, Bs @ map normt As, C) |
|
250 | 1203 |
else if match then raise Bicompose |
1204 |
else (*normalize the new rule fully*) |
|
1205 |
(ntps, map normt (Bs @ As), normt C) |
|
1206 |
end |
|
1258 | 1207 |
val th = (*tuned fix_shyps*) |
1208 |
Thm{sign=sign, |
|
1209 |
shyps=add_env_sorts (env, rshyps union sshyps), |
|
1210 |
hyps=rhyps union shyps, |
|
1211 |
maxidx=maxidx, prop= Logic.rule_of normp} |
|
0 | 1212 |
in cons(th, thq) end handle Bicompose => thq |
1213 |
val (rtpairs,rhorn) = Logic.strip_flexpairs(rprop); |
|
1214 |
val (rAs,B) = Logic.strip_prems(nsubgoal, [], rhorn) |
|
1215 |
handle TERM _ => raise THM("bicompose: rule", 0, [orule,state]); |
|
1216 |
(*Modify assumptions, deleting n-th if n>0 for e-resolution*) |
|
1217 |
fun newAs(As0, n, dpairs, tpairs) = |
|
1218 |
let val As1 = if !Logic.auto_rename orelse not lifted then As0 |
|
250 | 1219 |
else map (rename_bvars(dpairs,tpairs,B)) As0 |
0 | 1220 |
in (map (Logic.flatten_params n) As1) |
250 | 1221 |
handle TERM _ => |
1222 |
raise THM("bicompose: 1st premise", 0, [orule]) |
|
0 | 1223 |
end; |
1224 |
val env = Envir.empty(max[rmax,smax]); |
|
1225 |
val BBi = if lifted then strip_assums2(B,Bi) else (B,Bi); |
|
1226 |
val dpairs = BBi :: (rtpairs@stpairs); |
|
1227 |
(*elim-resolution: try each assumption in turn. Initially n=1*) |
|
1228 |
fun tryasms (_, _, []) = null |
|
1229 |
| tryasms (As, n, (t,u)::apairs) = |
|
250 | 1230 |
(case pull(Unify.unifiers(sign, env, (t,u)::dpairs)) of |
1231 |
None => tryasms (As, n+1, apairs) |
|
1232 |
| cell as Some((_,tpairs),_) => |
|
1233 |
its_right (addth (newAs(As, n, [BBi,(u,t)], tpairs))) |
|
1234 |
(seqof (fn()=> cell), |
|
1235 |
seqof (fn()=> pull (tryasms (As, n+1, apairs))))); |
|
0 | 1236 |
fun eres [] = raise THM("bicompose: no premises", 0, [orule,state]) |
1237 |
| eres (A1::As) = tryasms (As, 1, Logic.assum_pairs A1); |
|
1238 |
(*ordinary resolution*) |
|
1239 |
fun res(None) = null |
|
250 | 1240 |
| res(cell as Some((_,tpairs),_)) = |
1241 |
its_right (addth(newAs(rev rAs, 0, [BBi], tpairs))) |
|
1242 |
(seqof (fn()=> cell), null) |
|
0 | 1243 |
in if eres_flg then eres(rev rAs) |
1244 |
else res(pull(Unify.unifiers(sign, env, dpairs))) |
|
1245 |
end; |
|
1246 |
end; (*open Sequence*) |
|
1247 |
||
1248 |
||
1249 |
fun bicompose match arg i state = |
|
1250 |
bicompose_aux match (state, dest_state(state,i), false) arg; |
|
1251 |
||
1252 |
(*Quick test whether rule is resolvable with the subgoal with hyps Hs |
|
1253 |
and conclusion B. If eres_flg then checks 1st premise of rule also*) |
|
1254 |
fun could_bires (Hs, B, eres_flg, rule) = |
|
1255 |
let fun could_reshyp (A1::_) = exists (apl(A1,could_unify)) Hs |
|
250 | 1256 |
| could_reshyp [] = false; (*no premise -- illegal*) |
1257 |
in could_unify(concl_of rule, B) andalso |
|
1258 |
(not eres_flg orelse could_reshyp (prems_of rule)) |
|
0 | 1259 |
end; |
1260 |
||
1261 |
(*Bi-resolution of a state with a list of (flag,rule) pairs. |
|
1262 |
Puts the rule above: rule/state. Renames vars in the rules. *) |
|
250 | 1263 |
fun biresolution match brules i state = |
0 | 1264 |
let val lift = lift_rule(state, i); |
250 | 1265 |
val (stpairs, Bs, Bi, C) = dest_state(state,i) |
1266 |
val B = Logic.strip_assums_concl Bi; |
|
1267 |
val Hs = Logic.strip_assums_hyp Bi; |
|
1268 |
val comp = bicompose_aux match (state, (stpairs, Bs, Bi, C), true); |
|
1269 |
fun res [] = Sequence.null |
|
1270 |
| res ((eres_flg, rule)::brules) = |
|
1271 |
if could_bires (Hs, B, eres_flg, rule) |
|
1160 | 1272 |
then Sequence.seqof (*delay processing remainder till needed*) |
250 | 1273 |
(fn()=> Some(comp (eres_flg, lift rule, nprems_of rule), |
1274 |
res brules)) |
|
1275 |
else res brules |
|
0 | 1276 |
in Sequence.flats (res brules) end; |
1277 |
||
1278 |
||
1279 |
||
1280 |
(*** Meta simp sets ***) |
|
1281 |
||
288
b00ce6a1fe27
Implemented "ordered rewriting": rules which merely permute variables, such
nipkow
parents:
274
diff
changeset
|
1282 |
type rrule = {thm:thm, lhs:term, perm:bool}; |
b00ce6a1fe27
Implemented "ordered rewriting": rules which merely permute variables, such
nipkow
parents:
274
diff
changeset
|
1283 |
type cong = {thm:thm, lhs:term}; |
0 | 1284 |
datatype meta_simpset = |
405
c97514f63633
Internale optimization of the simplifier: in case a subterm stays unchanged,
nipkow
parents:
399
diff
changeset
|
1285 |
Mss of {net:rrule Net.net, congs:(string * cong)list, bounds:string list, |
0 | 1286 |
prems: thm list, mk_rews: thm -> thm list}; |
1287 |
||
1288 |
(*A "mss" contains data needed during conversion: |
|
1289 |
net: discrimination net of rewrite rules |
|
1290 |
congs: association list of congruence rules |
|
405
c97514f63633
Internale optimization of the simplifier: in case a subterm stays unchanged,
nipkow
parents:
399
diff
changeset
|
1291 |
bounds: names of bound variables already used; |
c97514f63633
Internale optimization of the simplifier: in case a subterm stays unchanged,
nipkow
parents:
399
diff
changeset
|
1292 |
for generating new names when rewriting under lambda abstractions |
0 | 1293 |
mk_rews: used when local assumptions are added |
1294 |
*) |
|
1295 |
||
405
c97514f63633
Internale optimization of the simplifier: in case a subterm stays unchanged,
nipkow
parents:
399
diff
changeset
|
1296 |
val empty_mss = Mss{net= Net.empty, congs= [], bounds=[], prems= [], |
0 | 1297 |
mk_rews = K[]}; |
1298 |
||
1299 |
exception SIMPLIFIER of string * thm; |
|
1300 |
||
229
4002c4cd450c
Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents:
225
diff
changeset
|
1301 |
fun prtm a sign t = (writeln a; writeln(Sign.string_of_term sign t)); |
0 | 1302 |
|
209 | 1303 |
val trace_simp = ref false; |
1304 |
||
229
4002c4cd450c
Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents:
225
diff
changeset
|
1305 |
fun trace_term a sign t = if !trace_simp then prtm a sign t else (); |
209 | 1306 |
|
1307 |
fun trace_thm a (Thm{sign,prop,...}) = trace_term a sign prop; |
|
1308 |
||
427
4ce2ce940faf
ordered rewriting applies to conditional rules as well now
nipkow
parents:
421
diff
changeset
|
1309 |
fun vperm(Var _, Var _) = true |
4ce2ce940faf
ordered rewriting applies to conditional rules as well now
nipkow
parents:
421
diff
changeset
|
1310 |
| vperm(Abs(_,_,s), Abs(_,_,t)) = vperm(s,t) |
4ce2ce940faf
ordered rewriting applies to conditional rules as well now
nipkow
parents:
421
diff
changeset
|
1311 |
| vperm(t1$t2, u1$u2) = vperm(t1,u1) andalso vperm(t2,u2) |
4ce2ce940faf
ordered rewriting applies to conditional rules as well now
nipkow
parents:
421
diff
changeset
|
1312 |
| vperm(t,u) = (t=u); |
288
b00ce6a1fe27
Implemented "ordered rewriting": rules which merely permute variables, such
nipkow
parents:
274
diff
changeset
|
1313 |
|
427
4ce2ce940faf
ordered rewriting applies to conditional rules as well now
nipkow
parents:
421
diff
changeset
|
1314 |
fun var_perm(t,u) = vperm(t,u) andalso |
4ce2ce940faf
ordered rewriting applies to conditional rules as well now
nipkow
parents:
421
diff
changeset
|
1315 |
eq_set(add_term_vars(t,[]), add_term_vars(u,[])) |
288
b00ce6a1fe27
Implemented "ordered rewriting": rules which merely permute variables, such
nipkow
parents:
274
diff
changeset
|
1316 |
|
0 | 1317 |
(*simple test for looping rewrite*) |
1318 |
fun loops sign prems (lhs,rhs) = |
|
1023
f5f36bdc8003
Fixed bug in the simplifier: added uses of maxidx_of_term to make sure that
nipkow
parents:
979
diff
changeset
|
1319 |
is_Var(lhs) |
f5f36bdc8003
Fixed bug in the simplifier: added uses of maxidx_of_term to make sure that
nipkow
parents:
979
diff
changeset
|
1320 |
orelse |
f5f36bdc8003
Fixed bug in the simplifier: added uses of maxidx_of_term to make sure that
nipkow
parents:
979
diff
changeset
|
1321 |
(exists (apl(lhs, Logic.occs)) (rhs::prems)) |
f5f36bdc8003
Fixed bug in the simplifier: added uses of maxidx_of_term to make sure that
nipkow
parents:
979
diff
changeset
|
1322 |
orelse |
f5f36bdc8003
Fixed bug in the simplifier: added uses of maxidx_of_term to make sure that
nipkow
parents:
979
diff
changeset
|
1323 |
(null(prems) andalso |
f5f36bdc8003
Fixed bug in the simplifier: added uses of maxidx_of_term to make sure that
nipkow
parents:
979
diff
changeset
|
1324 |
Pattern.matches (#tsig(Sign.rep_sg sign)) (lhs,rhs)); |
1028 | 1325 |
(* the condition "null(prems)" in the last case is necessary because |
1326 |
conditional rewrites with extra variables in the conditions may terminate |
|
1327 |
although the rhs is an instance of the lhs. Example: |
|
1328 |
?m < ?n ==> f(?n) == f(?m) |
|
1329 |
*) |
|
0 | 1330 |
|
1238 | 1331 |
fun mk_rrule raw_thm = |
1332 |
let |
|
1258 | 1333 |
val thm = strip_shyps raw_thm; |
1238 | 1334 |
val Thm{sign,prop,maxidx,...} = thm; |
1335 |
val prems = Logic.strip_imp_prems prop |
|
678
6151b7f3b606
Modified pattern.ML to perform proper matching of Higher-Order Patterns.
nipkow
parents:
659
diff
changeset
|
1336 |
val concl = Logic.strip_imp_concl prop |
6151b7f3b606
Modified pattern.ML to perform proper matching of Higher-Order Patterns.
nipkow
parents:
659
diff
changeset
|
1337 |
val (lhs,_) = Logic.dest_equals concl handle TERM _ => |
0 | 1338 |
raise SIMPLIFIER("Rewrite rule not a meta-equality",thm) |
678
6151b7f3b606
Modified pattern.ML to perform proper matching of Higher-Order Patterns.
nipkow
parents:
659
diff
changeset
|
1339 |
val econcl = Pattern.eta_contract concl |
6151b7f3b606
Modified pattern.ML to perform proper matching of Higher-Order Patterns.
nipkow
parents:
659
diff
changeset
|
1340 |
val (elhs,erhs) = Logic.dest_equals econcl |
6151b7f3b606
Modified pattern.ML to perform proper matching of Higher-Order Patterns.
nipkow
parents:
659
diff
changeset
|
1341 |
val perm = var_perm(elhs,erhs) andalso not(elhs aconv erhs) |
6151b7f3b606
Modified pattern.ML to perform proper matching of Higher-Order Patterns.
nipkow
parents:
659
diff
changeset
|
1342 |
andalso not(is_Var(elhs)) |
1220 | 1343 |
in |
1258 | 1344 |
if not perm andalso loops sign prems (elhs,erhs) then |
1220 | 1345 |
(prtm "Warning: ignoring looping rewrite rule" sign prop; None) |
288
b00ce6a1fe27
Implemented "ordered rewriting": rules which merely permute variables, such
nipkow
parents:
274
diff
changeset
|
1346 |
else Some{thm=thm,lhs=lhs,perm=perm} |
0 | 1347 |
end; |
1348 |
||
87 | 1349 |
local |
1350 |
fun eq({thm=Thm{prop=p1,...},...}:rrule, |
|
1351 |
{thm=Thm{prop=p2,...},...}:rrule) = p1 aconv p2 |
|
1352 |
in |
|
1353 |
||
405
c97514f63633
Internale optimization of the simplifier: in case a subterm stays unchanged,
nipkow
parents:
399
diff
changeset
|
1354 |
fun add_simp(mss as Mss{net,congs,bounds,prems,mk_rews}, |
0 | 1355 |
thm as Thm{sign,prop,...}) = |
87 | 1356 |
case mk_rrule thm of |
1357 |
None => mss |
|
1358 |
| Some(rrule as {lhs,...}) => |
|
209 | 1359 |
(trace_thm "Adding rewrite rule:" thm; |
1360 |
Mss{net= (Net.insert_term((lhs,rrule),net,eq) |
|
1361 |
handle Net.INSERT => |
|
87 | 1362 |
(prtm "Warning: ignoring duplicate rewrite rule" sign prop; |
1363 |
net)), |
|
405
c97514f63633
Internale optimization of the simplifier: in case a subterm stays unchanged,
nipkow
parents:
399
diff
changeset
|
1364 |
congs=congs, bounds=bounds, prems=prems,mk_rews=mk_rews}); |
87 | 1365 |
|
405
c97514f63633
Internale optimization of the simplifier: in case a subterm stays unchanged,
nipkow
parents:
399
diff
changeset
|
1366 |
fun del_simp(mss as Mss{net,congs,bounds,prems,mk_rews}, |
87 | 1367 |
thm as Thm{sign,prop,...}) = |
1368 |
case mk_rrule thm of |
|
1369 |
None => mss |
|
1370 |
| Some(rrule as {lhs,...}) => |
|
1371 |
Mss{net= (Net.delete_term((lhs,rrule),net,eq) |
|
1372 |
handle Net.INSERT => |
|
1373 |
(prtm "Warning: rewrite rule not in simpset" sign prop; |
|
1374 |
net)), |
|
405
c97514f63633
Internale optimization of the simplifier: in case a subterm stays unchanged,
nipkow
parents:
399
diff
changeset
|
1375 |
congs=congs, bounds=bounds, prems=prems,mk_rews=mk_rews} |
87 | 1376 |
|
1377 |
end; |
|
0 | 1378 |
|
1379 |
val add_simps = foldl add_simp; |
|
87 | 1380 |
val del_simps = foldl del_simp; |
0 | 1381 |
|