| author | wenzelm | 
| Tue, 28 Oct 1997 17:37:46 +0100 | |
| changeset 4023 | a9dc0484c903 | 
| parent 4018 | 09b77900af0f | 
| child 4036 | bd686e39bff8 | 
| 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  | 
| 1529 | 7  | 
theorems, meta rules (including resolution and simplification).  | 
| 0 | 8  | 
*)  | 
9  | 
||
| 250 | 10  | 
signature THM =  | 
| 1503 | 11  | 
sig  | 
| 1160 | 12  | 
(*certified types*)  | 
| 
387
 
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
 
wenzelm 
parents: 
309 
diff
changeset
 | 
13  | 
type ctyp  | 
| 1238 | 14  | 
  val rep_ctyp          : ctyp -> {sign: Sign.sg, T: typ}
 | 
15  | 
val typ_of : ctyp -> typ  | 
|
16  | 
val ctyp_of : Sign.sg -> typ -> ctyp  | 
|
17  | 
val read_ctyp : Sign.sg -> string -> ctyp  | 
|
| 1160 | 18  | 
|
19  | 
(*certified terms*)  | 
|
20  | 
type cterm  | 
|
| 
1493
 
e936723cb94d
added dest_comb, dest_abs and mk_prop for manipulating cterms
 
clasohm 
parents: 
1460 
diff
changeset
 | 
21  | 
exception CTERM of string  | 
| 
 
e936723cb94d
added dest_comb, dest_abs and mk_prop for manipulating cterms
 
clasohm 
parents: 
1460 
diff
changeset
 | 
22  | 
  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
 | 
23  | 
maxidx: int}  | 
| 1238 | 24  | 
val term_of : cterm -> term  | 
25  | 
val cterm_of : Sign.sg -> term -> cterm  | 
|
| 2671 | 26  | 
val ctyp_of_term : cterm -> ctyp  | 
| 1238 | 27  | 
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
 | 
28  | 
val read_cterms : Sign.sg -> string list * typ list -> cterm list  | 
| 1238 | 29  | 
val cterm_fun : (term -> term) -> (cterm -> cterm)  | 
| 
1493
 
e936723cb94d
added dest_comb, dest_abs and mk_prop for manipulating cterms
 
clasohm 
parents: 
1460 
diff
changeset
 | 
30  | 
val dest_comb : cterm -> cterm * cterm  | 
| 
 
e936723cb94d
added dest_comb, dest_abs and mk_prop for manipulating cterms
 
clasohm 
parents: 
1460 
diff
changeset
 | 
31  | 
val dest_abs : cterm -> cterm * cterm  | 
| 
1703
 
e22ad43bab5f
moved dest_cimplies to drule.ML; added adjust_maxidx
 
clasohm 
parents: 
1659 
diff
changeset
 | 
32  | 
val adjust_maxidx : cterm -> cterm  | 
| 
1516
 
96286c4e32de
removed mk_prop; added capply; simplified dest_abs
 
clasohm 
parents: 
1503 
diff
changeset
 | 
33  | 
val capply : cterm -> cterm -> cterm  | 
| 1517 | 34  | 
val cabs : cterm -> cterm -> cterm  | 
| 1238 | 35  | 
val read_def_cterm :  | 
| 1160 | 36  | 
Sign.sg * (indexname -> typ option) * (indexname -> sort option) ->  | 
37  | 
string list -> bool -> string * typ -> cterm * (indexname * typ) list  | 
|
38  | 
||
| 2671 | 39  | 
(*proof terms [must DUPLICATE declaration as a specification]*)  | 
| 
1597
 
54ece585bf62
name_thm no longer takes a theory argument, as the
 
paulson 
parents: 
1580 
diff
changeset
 | 
40  | 
datatype deriv_kind = MinDeriv | ThmDeriv | FullDeriv;  | 
| 2386 | 41  | 
val keep_derivs : deriv_kind ref  | 
| 1529 | 42  | 
datatype rule =  | 
| 2386 | 43  | 
MinProof  | 
| 3812 | 44  | 
| Oracle of theory * string * Sign.sg * exn  | 
| 2671 | 45  | 
| Axiom of theory * string  | 
46  | 
| Theorem of string  | 
|
47  | 
| Assume of cterm  | 
|
48  | 
| Implies_intr of cterm  | 
|
| 1529 | 49  | 
| Implies_intr_shyps  | 
50  | 
| Implies_intr_hyps  | 
|
51  | 
| Implies_elim  | 
|
| 2671 | 52  | 
| Forall_intr of cterm  | 
53  | 
| Forall_elim of cterm  | 
|
54  | 
| Reflexive of cterm  | 
|
| 1529 | 55  | 
| Symmetric  | 
56  | 
| Transitive  | 
|
| 2671 | 57  | 
| Beta_conversion of cterm  | 
| 1529 | 58  | 
| Extensional  | 
| 2671 | 59  | 
| Abstract_rule of string * cterm  | 
| 1529 | 60  | 
| Combination  | 
61  | 
| Equal_intr  | 
|
62  | 
| Equal_elim  | 
|
| 2671 | 63  | 
| Trivial of cterm  | 
64  | 
| Lift_rule of cterm * int  | 
|
65  | 
| Assumption of int * Envir.env option  | 
|
66  | 
| Rotate_rule of int * int  | 
|
67  | 
| Instantiate of (indexname * ctyp) list * (cterm * cterm) list  | 
|
68  | 
| Bicompose of bool * bool * int * int * Envir.env  | 
|
69  | 
| Flexflex_rule of Envir.env  | 
|
70  | 
| Class_triv of theory * class  | 
|
| 1529 | 71  | 
| VarifyT  | 
72  | 
| FreezeT  | 
|
| 2671 | 73  | 
| RewriteC of cterm  | 
74  | 
| CongC of cterm  | 
|
75  | 
| Rewrite_cterm of cterm  | 
|
76  | 
| Rename_params_rule of string list * int;  | 
|
| 1529 | 77  | 
|
| 
1597
 
54ece585bf62
name_thm no longer takes a theory argument, as the
 
paulson 
parents: 
1580 
diff
changeset
 | 
78  | 
type deriv (* = rule mtree *)  | 
| 1529 | 79  | 
|
| 1160 | 80  | 
(*meta theorems*)  | 
81  | 
type thm  | 
|
82  | 
exception THM of string * int * thm list  | 
|
| 1529 | 83  | 
  val rep_thm           : thm -> {sign: Sign.sg, der: deriv, maxidx: int,
 | 
| 2386 | 84  | 
shyps: sort list, hyps: term list,  | 
85  | 
prop: term}  | 
|
| 1529 | 86  | 
  val crep_thm          : thm -> {sign: Sign.sg, der: deriv, maxidx: int,
 | 
| 2386 | 87  | 
shyps: sort list, hyps: cterm list,  | 
88  | 
prop: cterm}  | 
|
| 3994 | 89  | 
val eq_thm : thm * thm -> bool  | 
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
90  | 
val sign_of_thm : thm -> Sign.sg  | 
| 3895 | 91  | 
val transfer : theory -> thm -> thm  | 
| 1238 | 92  | 
val tpairs_of : thm -> (term * term) list  | 
93  | 
val prems_of : thm -> term list  | 
|
94  | 
val nprems_of : thm -> int  | 
|
95  | 
val concl_of : thm -> term  | 
|
96  | 
val cprop_of : thm -> cterm  | 
|
97  | 
val extra_shyps : thm -> sort list  | 
|
| 3061 | 98  | 
val force_strip_shyps : bool ref (* FIXME tmp (since 1995/08/01) *)  | 
| 1238 | 99  | 
val strip_shyps : thm -> thm  | 
100  | 
val implies_intr_shyps: thm -> thm  | 
|
| 3812 | 101  | 
val get_axiom : theory -> xstring -> thm  | 
| 
1597
 
54ece585bf62
name_thm no longer takes a theory argument, as the
 
paulson 
parents: 
1580 
diff
changeset
 | 
102  | 
val name_thm : string * thm -> thm  | 
| 4018 | 103  | 
val name_of_thm : thm -> string  | 
| 1238 | 104  | 
val axioms_of : theory -> (string * thm) list  | 
| 1160 | 105  | 
|
106  | 
(*meta rules*)  | 
|
| 1238 | 107  | 
val assume : cterm -> thm  | 
| 1416 | 108  | 
val compress : thm -> thm  | 
| 1238 | 109  | 
val implies_intr : cterm -> thm -> thm  | 
110  | 
val implies_elim : thm -> thm -> thm  | 
|
111  | 
val forall_intr : cterm -> thm -> thm  | 
|
112  | 
val forall_elim : cterm -> thm -> 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  | 
|
| 2671 | 135  | 
val rotate_rule : int -> int -> thm -> thm  | 
| 1160 | 136  | 
val rename_params_rule: string list * int -> thm -> thm  | 
| 1238 | 137  | 
val bicompose : bool -> bool * thm * int ->  | 
| 1160 | 138  | 
int -> thm -> thm Sequence.seq  | 
| 1238 | 139  | 
val biresolution : bool -> (bool * thm) list ->  | 
| 1160 | 140  | 
int -> thm -> thm Sequence.seq  | 
141  | 
||
142  | 
(*meta simplification*)  | 
|
| 3550 | 143  | 
exception SIMPLIFIER of string * thm  | 
| 1160 | 144  | 
type meta_simpset  | 
| 3550 | 145  | 
val dest_mss : meta_simpset ->  | 
146  | 
    {simps: thm list, congs: thm list, procs: (string * cterm list) list}
 | 
|
| 1238 | 147  | 
val empty_mss : meta_simpset  | 
| 3550 | 148  | 
val merge_mss : meta_simpset * meta_simpset -> meta_simpset  | 
| 1238 | 149  | 
val add_simps : meta_simpset * thm list -> meta_simpset  | 
150  | 
val del_simps : meta_simpset * thm list -> meta_simpset  | 
|
151  | 
val mss_of : thm list -> meta_simpset  | 
|
152  | 
val add_congs : meta_simpset * thm list -> meta_simpset  | 
|
| 2626 | 153  | 
val del_congs : meta_simpset * thm list -> meta_simpset  | 
| 2509 | 154  | 
val add_simprocs : meta_simpset *  | 
| 
3577
 
9715b6e3ec5f
added prems argument to simplification procedures;
 
wenzelm 
parents: 
3565 
diff
changeset
 | 
155  | 
(string * cterm list * (Sign.sg -> thm list -> term -> thm option) * stamp) list  | 
| 
 
9715b6e3ec5f
added prems argument to simplification procedures;
 
wenzelm 
parents: 
3565 
diff
changeset
 | 
156  | 
-> meta_simpset  | 
| 2509 | 157  | 
val del_simprocs : meta_simpset *  | 
| 
3577
 
9715b6e3ec5f
added prems argument to simplification procedures;
 
wenzelm 
parents: 
3565 
diff
changeset
 | 
158  | 
(string * cterm list * (Sign.sg -> thm list -> term -> thm option) * stamp) list  | 
| 
 
9715b6e3ec5f
added prems argument to simplification procedures;
 
wenzelm 
parents: 
3565 
diff
changeset
 | 
159  | 
-> meta_simpset  | 
| 1238 | 160  | 
val add_prems : meta_simpset * thm list -> meta_simpset  | 
161  | 
val prems_of_mss : meta_simpset -> thm list  | 
|
162  | 
val set_mk_rews : meta_simpset * (thm -> thm list) -> meta_simpset  | 
|
163  | 
val mk_rews_of_mss : meta_simpset -> thm -> thm list  | 
|
| 2509 | 164  | 
val set_termless : meta_simpset * (term * term -> bool) -> meta_simpset  | 
| 1238 | 165  | 
val trace_simp : bool ref  | 
166  | 
val rewrite_cterm : bool * bool -> meta_simpset ->  | 
|
| 1529 | 167  | 
(meta_simpset -> thm -> thm option) -> cterm -> thm  | 
| 1539 | 168  | 
|
| 3812 | 169  | 
val invoke_oracle : theory -> xstring -> Sign.sg * exn -> thm  | 
| 250 | 170  | 
end;  | 
| 0 | 171  | 
|
| 3550 | 172  | 
structure Thm: THM =  | 
| 0 | 173  | 
struct  | 
| 250 | 174  | 
|
| 
387
 
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
 
wenzelm 
parents: 
309 
diff
changeset
 | 
175  | 
(*** Certified terms and types ***)  | 
| 
 
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
 
wenzelm 
parents: 
309 
diff
changeset
 | 
176  | 
|
| 250 | 177  | 
(** certified types **)  | 
178  | 
||
179  | 
(*certified typs under a signature*)  | 
|
180  | 
||
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
181  | 
datatype ctyp = Ctyp of {sign_ref: Sign.sg_ref, T: typ};
 | 
| 250 | 182  | 
|
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
183  | 
fun rep_ctyp (Ctyp {sign_ref, T}) = {sign = Sign.deref sign_ref, T = T};
 | 
| 250 | 184  | 
fun typ_of (Ctyp {T, ...}) = T;
 | 
185  | 
||
186  | 
fun ctyp_of sign T =  | 
|
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
187  | 
  Ctyp {sign_ref = Sign.self_ref sign, T = Sign.certify_typ sign T};
 | 
| 250 | 188  | 
|
189  | 
fun read_ctyp sign s =  | 
|
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
190  | 
  Ctyp {sign_ref = Sign.self_ref 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
 | 
191  | 
|
| 
 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 
lcp 
parents: 
225 
diff
changeset
 | 
192  | 
|
| 
 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 
lcp 
parents: 
225 
diff
changeset
 | 
193  | 
|
| 250 | 194  | 
(** certified terms **)  | 
| 
229
 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 
lcp 
parents: 
225 
diff
changeset
 | 
195  | 
|
| 250 | 196  | 
(*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
 | 
197  | 
|
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
198  | 
datatype cterm = Cterm of {sign_ref: Sign.sg_ref, 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
 | 
199  | 
|
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
200  | 
fun rep_cterm (Cterm {sign_ref, t, T, maxidx}) =
 | 
| 
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
201  | 
  {sign = Sign.deref sign_ref, t = t, T = T, maxidx = maxidx};
 | 
| 
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
202  | 
|
| 250 | 203  | 
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
 | 
204  | 
|
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
205  | 
fun ctyp_of_term (Cterm {sign_ref, T, ...}) = Ctyp {sign_ref = sign_ref, T = T};
 | 
| 2671 | 206  | 
|
| 250 | 207  | 
(*create a cterm by checking a "raw" term with respect to a signature*)  | 
208  | 
fun cterm_of sign tm =  | 
|
209  | 
let val (t, T, maxidx) = Sign.certify_term sign tm  | 
|
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
210  | 
  in  Cterm {sign_ref = Sign.self_ref sign, t = t, T = T, maxidx = maxidx}
 | 
| 
1394
 
a1d2735f5ade
New function read_cterms is a combination of read_def_cterm and
 
paulson 
parents: 
1258 
diff
changeset
 | 
211  | 
end;  | 
| 
229
 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 
lcp 
parents: 
225 
diff
changeset
 | 
212  | 
|
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
213  | 
fun cterm_fun f (Cterm {sign_ref, t, ...}) = cterm_of (Sign.deref sign_ref) (f t);
 | 
| 250 | 214  | 
|
| 
229
 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 
lcp 
parents: 
225 
diff
changeset
 | 
215  | 
|
| 
1493
 
e936723cb94d
added dest_comb, dest_abs and mk_prop for manipulating cterms
 
clasohm 
parents: 
1460 
diff
changeset
 | 
216  | 
exception CTERM of string;  | 
| 
 
e936723cb94d
added dest_comb, dest_abs and mk_prop for manipulating cterms
 
clasohm 
parents: 
1460 
diff
changeset
 | 
217  | 
|
| 
 
e936723cb94d
added dest_comb, dest_abs and mk_prop for manipulating cterms
 
clasohm 
parents: 
1460 
diff
changeset
 | 
218  | 
(*Destruct application in cterms*)  | 
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
219  | 
fun dest_comb (Cterm {sign_ref, T, maxidx, t = A $ B}) =
 | 
| 
1493
 
e936723cb94d
added dest_comb, dest_abs and mk_prop for manipulating cterms
 
clasohm 
parents: 
1460 
diff
changeset
 | 
220  | 
let val typeA = fastype_of A;  | 
| 
 
e936723cb94d
added dest_comb, dest_abs and mk_prop for manipulating cterms
 
clasohm 
parents: 
1460 
diff
changeset
 | 
221  | 
val typeB =  | 
| 
 
e936723cb94d
added dest_comb, dest_abs and mk_prop for manipulating cterms
 
clasohm 
parents: 
1460 
diff
changeset
 | 
222  | 
            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
 | 
223  | 
| _ => error "Function type expected in dest_comb";  | 
| 
 
e936723cb94d
added dest_comb, dest_abs and mk_prop for manipulating cterms
 
clasohm 
parents: 
1460 
diff
changeset
 | 
224  | 
in  | 
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
225  | 
      (Cterm {sign_ref=sign_ref, maxidx=maxidx, t=A, T=typeA},
 | 
| 
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
226  | 
       Cterm {sign_ref=sign_ref, maxidx=maxidx, t=B, T=typeB})
 | 
| 
1493
 
e936723cb94d
added dest_comb, dest_abs and mk_prop for manipulating cterms
 
clasohm 
parents: 
1460 
diff
changeset
 | 
227  | 
end  | 
| 
 
e936723cb94d
added dest_comb, dest_abs and mk_prop for manipulating cterms
 
clasohm 
parents: 
1460 
diff
changeset
 | 
228  | 
| dest_comb _ = raise CTERM "dest_comb";  | 
| 
 
e936723cb94d
added dest_comb, dest_abs and mk_prop for manipulating cterms
 
clasohm 
parents: 
1460 
diff
changeset
 | 
229  | 
|
| 
 
e936723cb94d
added dest_comb, dest_abs and mk_prop for manipulating cterms
 
clasohm 
parents: 
1460 
diff
changeset
 | 
230  | 
(*Destruct abstraction in cterms*)  | 
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
231  | 
fun dest_abs (Cterm {sign_ref, T as Type("fun",[_,S]), maxidx, t=Abs(x,ty,M)}) = 
 | 
| 
1516
 
96286c4e32de
removed mk_prop; added capply; simplified dest_abs
 
clasohm 
parents: 
1503 
diff
changeset
 | 
232  | 
let val (y,N) = variant_abs (x,ty,M)  | 
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
233  | 
      in (Cterm {sign_ref = sign_ref, T = ty, maxidx = 0, t = Free(y,ty)},
 | 
| 
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
234  | 
          Cterm {sign_ref = sign_ref, T = S, maxidx = maxidx, t = N})
 | 
| 
1493
 
e936723cb94d
added dest_comb, dest_abs and mk_prop for manipulating cterms
 
clasohm 
parents: 
1460 
diff
changeset
 | 
235  | 
end  | 
| 
 
e936723cb94d
added dest_comb, dest_abs and mk_prop for manipulating cterms
 
clasohm 
parents: 
1460 
diff
changeset
 | 
236  | 
| dest_abs _ = raise CTERM "dest_abs";  | 
| 
 
e936723cb94d
added dest_comb, dest_abs and mk_prop for manipulating cterms
 
clasohm 
parents: 
1460 
diff
changeset
 | 
237  | 
|
| 2147 | 238  | 
(*Makes maxidx precise: it is often too big*)  | 
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
239  | 
fun adjust_maxidx (ct as Cterm {sign_ref, T, t, maxidx, ...}) =
 | 
| 2147 | 240  | 
if maxidx = ~1 then ct  | 
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
241  | 
  else  Cterm {sign_ref = sign_ref, T = T, maxidx = maxidx_of_term t, t = t};
 | 
| 
1703
 
e22ad43bab5f
moved dest_cimplies to drule.ML; added adjust_maxidx
 
clasohm 
parents: 
1659 
diff
changeset
 | 
242  | 
|
| 
1516
 
96286c4e32de
removed mk_prop; added capply; simplified dest_abs
 
clasohm 
parents: 
1503 
diff
changeset
 | 
243  | 
(*Form cterm out of a function and an argument*)  | 
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
244  | 
fun capply (Cterm {t=f, sign_ref=sign_ref1, T=Type("fun",[dty,rty]), maxidx=maxidx1})
 | 
| 
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
245  | 
           (Cterm {t=x, sign_ref=sign_ref2, T, maxidx=maxidx2}) =
 | 
| 
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
246  | 
      if T = dty then Cterm{t=f$x, sign_ref=Sign.merge_refs(sign_ref1,sign_ref2), T=rty,
 | 
| 2147 | 247  | 
maxidx=Int.max(maxidx1, maxidx2)}  | 
| 
1516
 
96286c4e32de
removed mk_prop; added capply; simplified dest_abs
 
clasohm 
parents: 
1503 
diff
changeset
 | 
248  | 
else raise CTERM "capply: types don't agree"  | 
| 
 
96286c4e32de
removed mk_prop; added capply; simplified dest_abs
 
clasohm 
parents: 
1503 
diff
changeset
 | 
249  | 
| capply _ _ = raise CTERM "capply: first arg is not a function"  | 
| 250 | 250  | 
|
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
251  | 
fun cabs (Cterm {t=Free(a,ty), sign_ref=sign_ref1, T=T1, maxidx=maxidx1})
 | 
| 
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
252  | 
         (Cterm {t=t2, sign_ref=sign_ref2, T=T2, maxidx=maxidx2}) =
 | 
| 
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
253  | 
      Cterm {t=absfree(a,ty,t2), sign_ref=Sign.merge_refs(sign_ref1,sign_ref2),
 | 
| 2147 | 254  | 
T = ty --> T2, maxidx=Int.max(maxidx1, maxidx2)}  | 
| 1517 | 255  | 
| cabs _ _ = raise CTERM "cabs: first arg is not a free variable";  | 
| 
229
 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 
lcp 
parents: 
225 
diff
changeset
 | 
256  | 
|
| 2509 | 257  | 
|
258  | 
||
| 574 | 259  | 
(** read cterms **) (*exception ERROR*)  | 
| 250 | 260  | 
|
261  | 
(*read term, infer types, certify term*)  | 
|
| 
949
 
83c588d6fee9
Changed treatment of during type inference internally generated type
 
nipkow 
parents: 
922 
diff
changeset
 | 
262  | 
fun read_def_cterm (sign, types, sorts) used freeze (a, T) =  | 
| 250 | 263  | 
let  | 
| 574 | 264  | 
val T' = Sign.certify_typ sign T  | 
265  | 
handle TYPE (msg, _, _) => error msg;  | 
|
| 623 | 266  | 
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
 | 
267  | 
val (_, t', tye) =  | 
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
268  | 
Sign.infer_types sign types sorts used freeze (ts, T');  | 
| 574 | 269  | 
val ct = cterm_of sign t'  | 
| 2979 | 270  | 
handle TYPE (msg, _, _) => error msg  | 
| 2386 | 271  | 
| TERM (msg, _) => error msg;  | 
| 250 | 272  | 
in (ct, tye) end;  | 
| 
229
 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 
lcp 
parents: 
225 
diff
changeset
 | 
273  | 
|
| 
949
 
83c588d6fee9
Changed treatment of during type inference internally generated type
 
nipkow 
parents: 
922 
diff
changeset
 | 
274  | 
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
 | 
275  | 
|
| 
1394
 
a1d2735f5ade
New function read_cterms is a combination of read_def_cterm and
 
paulson 
parents: 
1258 
diff
changeset
 | 
276  | 
(*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
 | 
277  | 
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
 | 
278  | 
not practical.*)  | 
| 3789 | 279  | 
fun read_cterms sg (bs, Ts) =  | 
| 
1394
 
a1d2735f5ade
New function read_cterms is a combination of read_def_cterm and
 
paulson 
parents: 
1258 
diff
changeset
 | 
280  | 
let  | 
| 3789 | 281  | 
    val {tsig, syn, ...} = Sign.rep_sg sg;
 | 
| 2979 | 282  | 
fun read (b, T) =  | 
283  | 
(case Syntax.read syn T b of  | 
|
284  | 
[t] => t  | 
|
285  | 
      | _  => error ("Error or ambiguity in parsing of " ^ b));
 | 
|
286  | 
||
| 3789 | 287  | 
val prt = setmp Syntax.show_brackets true (Sign.pretty_term sg);  | 
288  | 
val prT = Sign.pretty_typ sg;  | 
|
| 2979 | 289  | 
val (us, _) =  | 
| 3789 | 290  | 
(* FIXME Sign.infer_types!? *)  | 
291  | 
Type.infer_types prt prT tsig (Sign.const_type sg) (K None) (K None)  | 
|
292  | 
(Sign.intern_const sg) (Sign.intern_tycons sg) (Sign.intern_sort sg)  | 
|
293  | 
[] true (map (Sign.certify_typ sg) Ts) (ListPair.map read (bs, Ts));  | 
|
294  | 
in map (cterm_of sg) us end  | 
|
| 2979 | 295  | 
handle TYPE (msg, _, _) => error msg  | 
| 
1394
 
a1d2735f5ade
New function read_cterms is a combination of read_def_cterm and
 
paulson 
parents: 
1258 
diff
changeset
 | 
296  | 
| TERM (msg, _) => error msg;  | 
| 
 
a1d2735f5ade
New function read_cterms is a combination of read_def_cterm and
 
paulson 
parents: 
1258 
diff
changeset
 | 
297  | 
|
| 250 | 298  | 
|
299  | 
||
| 1529 | 300  | 
(*** Derivations ***)  | 
301  | 
||
302  | 
(*Names of rules in derivations. Includes logically trivial rules, if  | 
|
303  | 
executed in ML.*)  | 
|
304  | 
datatype rule =  | 
|
| 2386 | 305  | 
MinProof (*for building minimal proof terms*)  | 
| 3812 | 306  | 
| Oracle of theory * string * Sign.sg * exn (*oracles*)  | 
| 1529 | 307  | 
(*Axioms/theorems*)  | 
| 2386 | 308  | 
| Axiom of theory * string  | 
309  | 
| Theorem of string  | 
|
| 1529 | 310  | 
(*primitive inferences and compound versions of them*)  | 
| 2386 | 311  | 
| Assume of cterm  | 
312  | 
| Implies_intr of cterm  | 
|
| 1529 | 313  | 
| Implies_intr_shyps  | 
314  | 
| Implies_intr_hyps  | 
|
315  | 
| Implies_elim  | 
|
| 2386 | 316  | 
| Forall_intr of cterm  | 
317  | 
| Forall_elim of cterm  | 
|
318  | 
| Reflexive of cterm  | 
|
| 1529 | 319  | 
| Symmetric  | 
320  | 
| Transitive  | 
|
| 2386 | 321  | 
| Beta_conversion of cterm  | 
| 1529 | 322  | 
| Extensional  | 
| 2386 | 323  | 
| Abstract_rule of string * cterm  | 
| 1529 | 324  | 
| Combination  | 
325  | 
| Equal_intr  | 
|
326  | 
| Equal_elim  | 
|
327  | 
(*derived rules for tactical proof*)  | 
|
| 2386 | 328  | 
| Trivial of cterm  | 
329  | 
(*For lift_rule, the proof state is not a premise.  | 
|
330  | 
Use cterm instead of thm to avoid mutual recursion.*)  | 
|
331  | 
| Lift_rule of cterm * int  | 
|
332  | 
| Assumption of int * Envir.env option (*includes eq_assumption*)  | 
|
| 2671 | 333  | 
| Rotate_rule of int * int  | 
| 2386 | 334  | 
| Instantiate of (indexname * ctyp) list * (cterm * cterm) list  | 
335  | 
| Bicompose of bool * bool * int * int * Envir.env  | 
|
336  | 
| Flexflex_rule of Envir.env (*identifies unifier chosen*)  | 
|
| 1529 | 337  | 
(*other derived rules*)  | 
| 2509 | 338  | 
| Class_triv of theory * class  | 
| 1529 | 339  | 
| VarifyT  | 
340  | 
| FreezeT  | 
|
341  | 
(*for the simplifier*)  | 
|
| 2386 | 342  | 
| RewriteC of cterm  | 
343  | 
| CongC of cterm  | 
|
344  | 
| Rewrite_cterm of cterm  | 
|
| 1529 | 345  | 
(*Logical identities, recorded since they are part of the proof process*)  | 
| 2386 | 346  | 
| Rename_params_rule of string list * int;  | 
| 1529 | 347  | 
|
348  | 
||
| 
1597
 
54ece585bf62
name_thm no longer takes a theory argument, as the
 
paulson 
parents: 
1580 
diff
changeset
 | 
349  | 
type deriv = rule mtree;  | 
| 1529 | 350  | 
|
| 
1597
 
54ece585bf62
name_thm no longer takes a theory argument, as the
 
paulson 
parents: 
1580 
diff
changeset
 | 
351  | 
datatype deriv_kind = MinDeriv | ThmDeriv | FullDeriv;  | 
| 1529 | 352  | 
|
| 
1597
 
54ece585bf62
name_thm no longer takes a theory argument, as the
 
paulson 
parents: 
1580 
diff
changeset
 | 
353  | 
val keep_derivs = ref MinDeriv;  | 
| 1529 | 354  | 
|
355  | 
||
| 
1597
 
54ece585bf62
name_thm no longer takes a theory argument, as the
 
paulson 
parents: 
1580 
diff
changeset
 | 
356  | 
(*Build a minimal derivation. Keep oracles; suppress atomic inferences;  | 
| 
 
54ece585bf62
name_thm no longer takes a theory argument, as the
 
paulson 
parents: 
1580 
diff
changeset
 | 
357  | 
retain Theorems or their underlying links; keep anything else*)  | 
| 
 
54ece585bf62
name_thm no longer takes a theory argument, as the
 
paulson 
parents: 
1580 
diff
changeset
 | 
358  | 
fun squash_derivs [] = []  | 
| 
 
54ece585bf62
name_thm no longer takes a theory argument, as the
 
paulson 
parents: 
1580 
diff
changeset
 | 
359  | 
| squash_derivs (der::ders) =  | 
| 
 
54ece585bf62
name_thm no longer takes a theory argument, as the
 
paulson 
parents: 
1580 
diff
changeset
 | 
360  | 
(case der of  | 
| 2386 | 361  | 
Join (Oracle _, _) => der :: squash_derivs ders  | 
362  | 
| Join (Theorem _, [der']) => if !keep_derivs=ThmDeriv  | 
|
363  | 
then der :: squash_derivs ders  | 
|
364  | 
else squash_derivs (der'::ders)  | 
|
365  | 
| Join (Axiom _, _) => if !keep_derivs=ThmDeriv  | 
|
366  | 
then der :: squash_derivs ders  | 
|
367  | 
else squash_derivs ders  | 
|
368  | 
| Join (_, []) => squash_derivs ders  | 
|
369  | 
| _ => der :: squash_derivs ders);  | 
|
| 
1597
 
54ece585bf62
name_thm no longer takes a theory argument, as the
 
paulson 
parents: 
1580 
diff
changeset
 | 
370  | 
|
| 1529 | 371  | 
|
372  | 
(*Ensure sharing of the most likely derivation, the empty one!*)  | 
|
| 
1597
 
54ece585bf62
name_thm no longer takes a theory argument, as the
 
paulson 
parents: 
1580 
diff
changeset
 | 
373  | 
val min_infer = Join (MinProof, []);  | 
| 1529 | 374  | 
|
375  | 
(*Make a minimal inference*)  | 
|
376  | 
fun make_min_infer [] = min_infer  | 
|
377  | 
| make_min_infer [der] = der  | 
|
| 
1597
 
54ece585bf62
name_thm no longer takes a theory argument, as the
 
paulson 
parents: 
1580 
diff
changeset
 | 
378  | 
| make_min_infer ders = Join (MinProof, ders);  | 
| 1529 | 379  | 
|
| 
1597
 
54ece585bf62
name_thm no longer takes a theory argument, as the
 
paulson 
parents: 
1580 
diff
changeset
 | 
380  | 
fun infer_derivs (rl, []) = Join (rl, [])  | 
| 1529 | 381  | 
| infer_derivs (rl, ders) =  | 
| 
1597
 
54ece585bf62
name_thm no longer takes a theory argument, as the
 
paulson 
parents: 
1580 
diff
changeset
 | 
382  | 
if !keep_derivs=FullDeriv then Join (rl, ders)  | 
| 1529 | 383  | 
else make_min_infer (squash_derivs ders);  | 
384  | 
||
385  | 
||
| 2509 | 386  | 
|
| 
387
 
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
 
wenzelm 
parents: 
309 
diff
changeset
 | 
387  | 
(*** Meta theorems ***)  | 
| 
229
 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 
lcp 
parents: 
225 
diff
changeset
 | 
388  | 
|
| 0 | 389  | 
datatype thm = Thm of  | 
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
390  | 
 {sign_ref: Sign.sg_ref,       (*mutable reference to signature*)
 | 
| 
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
391  | 
der: deriv, (*derivation*)  | 
| 
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
392  | 
maxidx: int, (*maximum index of any Var or TVar*)  | 
| 
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
393  | 
shyps: sort list, (*sort hypotheses*)  | 
| 
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
394  | 
hyps: term list, (*hypotheses*)  | 
| 
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
395  | 
prop: term}; (*conclusion*)  | 
| 0 | 396  | 
|
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
397  | 
fun rep_thm (Thm {sign_ref, der, maxidx, shyps, hyps, prop}) =
 | 
| 
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
398  | 
  {sign = Sign.deref sign_ref, der = der, maxidx = maxidx,
 | 
| 
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
399  | 
shyps = shyps, hyps = hyps, prop = prop};  | 
| 0 | 400  | 
|
| 1529 | 401  | 
(*Version of rep_thm returning cterms instead of terms*)  | 
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
402  | 
fun crep_thm (Thm {sign_ref, der, maxidx, shyps, hyps, prop}) =
 | 
| 
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
403  | 
  let fun ctermf max t = Cterm{sign_ref=sign_ref, t=t, T=propT, maxidx=max};
 | 
| 
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
404  | 
  in {sign = Sign.deref sign_ref, der = der, maxidx = maxidx, shyps = shyps,
 | 
| 1529 | 405  | 
hyps = map (ctermf ~1) hyps,  | 
406  | 
prop = ctermf maxidx prop}  | 
|
| 1517 | 407  | 
end;  | 
408  | 
||
| 
387
 
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
 
wenzelm 
parents: 
309 
diff
changeset
 | 
409  | 
(*errors involving theorems*)  | 
| 0 | 410  | 
exception THM of string * int * thm list;  | 
411  | 
||
| 3994 | 412  | 
(*equality of theorems uses equality of signatures and the  | 
413  | 
a-convertible test for terms*)  | 
|
414  | 
fun eq_thm (th1, th2) =  | 
|
415  | 
let  | 
|
416  | 
    val {sign = sg1, shyps = shyps1, hyps = hyps1, prop = prop1, ...} = rep_thm th1;
 | 
|
417  | 
    val {sign = sg2, shyps = shyps2, hyps = hyps2, prop = prop2, ...} = rep_thm th2;
 | 
|
418  | 
in  | 
|
419  | 
Sign.eq_sg (sg1, sg2) andalso  | 
|
420  | 
eq_set_sort (shyps1, shyps2) andalso  | 
|
421  | 
aconvs (hyps1, hyps2) andalso  | 
|
422  | 
prop1 aconv prop2  | 
|
423  | 
end;  | 
|
| 
387
 
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
 
wenzelm 
parents: 
309 
diff
changeset
 | 
424  | 
|
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
425  | 
fun sign_of_thm (Thm {sign_ref, ...}) = Sign.deref sign_ref;
 | 
| 0 | 426  | 
|
| 
387
 
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
 
wenzelm 
parents: 
309 
diff
changeset
 | 
427  | 
(*merge signatures of two theorems; raise exception if incompatible*)  | 
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
428  | 
fun merge_thm_sgs  | 
| 
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
429  | 
    (th1 as Thm {sign_ref = sgr1, ...}, th2 as Thm {sign_ref = sgr2, ...}) =
 | 
| 
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
430  | 
Sign.merge_refs (sgr1, sgr2) 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
 | 
431  | 
|
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
432  | 
(*transfer thm to super theory (non-destructive)*)  | 
| 3895 | 433  | 
fun transfer thy thm =  | 
434  | 
let  | 
|
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
435  | 
    val Thm {sign_ref, der, maxidx, shyps, hyps, prop} = thm;
 | 
| 
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
436  | 
val sign = Sign.deref sign_ref;  | 
| 3895 | 437  | 
val sign' = #sign (rep_theory thy);  | 
438  | 
in  | 
|
439  | 
if Sign.subsig (sign, sign') then  | 
|
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
440  | 
      Thm {sign_ref = Sign.self_ref sign', der = der, maxidx = maxidx,
 | 
| 3895 | 441  | 
shyps = shyps, hyps = hyps, prop = prop}  | 
442  | 
    else raise THM ("transfer: not a super theory", 0, [thm])
 | 
|
443  | 
end;  | 
|
| 
387
 
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
 
wenzelm 
parents: 
309 
diff
changeset
 | 
444  | 
|
| 
 
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
 
wenzelm 
parents: 
309 
diff
changeset
 | 
445  | 
(*maps object-rule to tpairs*)  | 
| 
 
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
 
wenzelm 
parents: 
309 
diff
changeset
 | 
446  | 
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
 | 
447  | 
|
| 
 
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
 
wenzelm 
parents: 
309 
diff
changeset
 | 
448  | 
(*maps object-rule to premises*)  | 
| 
 
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
 
wenzelm 
parents: 
309 
diff
changeset
 | 
449  | 
fun prems_of (Thm {prop, ...}) =
 | 
| 
 
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
 
wenzelm 
parents: 
309 
diff
changeset
 | 
450  | 
Logic.strip_imp_prems (Logic.skip_flexpairs prop);  | 
| 0 | 451  | 
|
452  | 
(*counts premises in a rule*)  | 
|
| 
387
 
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
 
wenzelm 
parents: 
309 
diff
changeset
 | 
453  | 
fun nprems_of (Thm {prop, ...}) =
 | 
| 
 
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
 
wenzelm 
parents: 
309 
diff
changeset
 | 
454  | 
Logic.count_prems (Logic.skip_flexpairs prop, 0);  | 
| 0 | 455  | 
|
| 
387
 
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
 
wenzelm 
parents: 
309 
diff
changeset
 | 
456  | 
(*maps object-rule to conclusion*)  | 
| 
 
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
 
wenzelm 
parents: 
309 
diff
changeset
 | 
457  | 
fun concl_of (Thm {prop, ...}) = Logic.strip_imp_concl prop;
 | 
| 0 | 458  | 
|
| 
387
 
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
 
wenzelm 
parents: 
309 
diff
changeset
 | 
459  | 
(*the statement of any thm is a cterm*)  | 
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
460  | 
fun cprop_of (Thm {sign_ref, maxidx, prop, ...}) =
 | 
| 
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
461  | 
  Cterm {sign_ref = sign_ref, 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
 | 
462  | 
|
| 
387
 
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
 
wenzelm 
parents: 
309 
diff
changeset
 | 
463  | 
|
| 0 | 464  | 
|
| 1238 | 465  | 
(** sort contexts of theorems **)  | 
466  | 
||
467  | 
(* basic utils *)  | 
|
468  | 
||
| 2163 | 469  | 
(*accumulate sorts suppressing duplicates; these are coded low levelly  | 
| 1238 | 470  | 
to improve efficiency a bit*)  | 
471  | 
||
472  | 
fun add_typ_sorts (Type (_, Ts), Ss) = add_typs_sorts (Ts, Ss)  | 
|
| 
2177
 
8b365a3a6ed1
Changed some mem, ins and union calls to be monomorphic
 
paulson 
parents: 
2163 
diff
changeset
 | 
473  | 
| add_typ_sorts (TFree (_, S), Ss) = ins_sort(S,Ss)  | 
| 
 
8b365a3a6ed1
Changed some mem, ins and union calls to be monomorphic
 
paulson 
parents: 
2163 
diff
changeset
 | 
474  | 
| add_typ_sorts (TVar (_, S), Ss) = ins_sort(S,Ss)  | 
| 1238 | 475  | 
and add_typs_sorts ([], Ss) = Ss  | 
476  | 
| add_typs_sorts (T :: Ts, Ss) = add_typs_sorts (Ts, add_typ_sorts (T, Ss));  | 
|
477  | 
||
478  | 
fun add_term_sorts (Const (_, T), Ss) = add_typ_sorts (T, Ss)  | 
|
479  | 
| add_term_sorts (Free (_, T), Ss) = add_typ_sorts (T, Ss)  | 
|
480  | 
| add_term_sorts (Var (_, T), Ss) = add_typ_sorts (T, Ss)  | 
|
481  | 
| add_term_sorts (Bound _, Ss) = Ss  | 
|
| 
2177
 
8b365a3a6ed1
Changed some mem, ins and union calls to be monomorphic
 
paulson 
parents: 
2163 
diff
changeset
 | 
482  | 
| add_term_sorts (Abs (_,T,t), Ss) = add_term_sorts (t, add_typ_sorts (T,Ss))  | 
| 1238 | 483  | 
| add_term_sorts (t $ u, Ss) = add_term_sorts (t, add_term_sorts (u, Ss));  | 
484  | 
||
485  | 
fun add_terms_sorts ([], Ss) = Ss  | 
|
| 
2177
 
8b365a3a6ed1
Changed some mem, ins and union calls to be monomorphic
 
paulson 
parents: 
2163 
diff
changeset
 | 
486  | 
| add_terms_sorts (t::ts, Ss) = add_terms_sorts (ts, add_term_sorts (t,Ss));  | 
| 1238 | 487  | 
|
| 1258 | 488  | 
fun env_codT (Envir.Envir {iTs, ...}) = map snd iTs;
 | 
489  | 
||
490  | 
fun add_env_sorts (env, Ss) =  | 
|
491  | 
add_terms_sorts (map snd (Envir.alist_of env),  | 
|
492  | 
add_typs_sorts (env_codT env, Ss));  | 
|
493  | 
||
| 1238 | 494  | 
fun add_thm_sorts (Thm {hyps, prop, ...}, Ss) =
 | 
495  | 
add_terms_sorts (hyps, add_term_sorts (prop, Ss));  | 
|
496  | 
||
497  | 
fun add_thms_shyps ([], Ss) = Ss  | 
|
498  | 
  | add_thms_shyps (Thm {shyps, ...} :: ths, Ss) =
 | 
|
| 
2177
 
8b365a3a6ed1
Changed some mem, ins and union calls to be monomorphic
 
paulson 
parents: 
2163 
diff
changeset
 | 
499  | 
add_thms_shyps (ths, union_sort(shyps,Ss));  | 
| 1238 | 500  | 
|
501  | 
||
502  | 
(*get 'dangling' sort constraints of a thm*)  | 
|
503  | 
fun extra_shyps (th as Thm {shyps, ...}) =
 | 
|
504  | 
shyps \\ add_thm_sorts (th, []);  | 
|
505  | 
||
506  | 
||
507  | 
(* fix_shyps *)  | 
|
508  | 
||
509  | 
(*preserve sort contexts of rule premises and substituted types*)  | 
|
510  | 
fun fix_shyps thms Ts thm =  | 
|
511  | 
let  | 
|
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
512  | 
    val Thm {sign_ref, der, maxidx, hyps, prop, ...} = thm;
 | 
| 1238 | 513  | 
val shyps =  | 
514  | 
add_thm_sorts (thm, add_typs_sorts (Ts, add_thms_shyps (thms, [])));  | 
|
515  | 
in  | 
|
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
516  | 
    Thm {sign_ref = sign_ref,
 | 
| 2386 | 517  | 
der = der, (*No new derivation, as other rules call this*)  | 
518  | 
maxidx = maxidx,  | 
|
519  | 
shyps = shyps, hyps = hyps, prop = prop}  | 
|
| 1238 | 520  | 
end;  | 
521  | 
||
522  | 
||
523  | 
(* strip_shyps *) (* FIXME improve? (e.g. only minimal extra sorts) *)  | 
|
524  | 
||
| 3061 | 525  | 
val force_strip_shyps = ref true; (* FIXME tmp (since 1995/08/01) *)  | 
| 1238 | 526  | 
|
527  | 
(*remove extra sorts that are known to be syntactically non-empty*)  | 
|
528  | 
fun strip_shyps thm =  | 
|
529  | 
let  | 
|
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
530  | 
    val Thm {sign_ref, der, maxidx, shyps, hyps, prop} = thm;
 | 
| 1238 | 531  | 
val sorts = add_thm_sorts (thm, []);  | 
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
532  | 
val maybe_empty = not o Sign.nonempty_sort (Sign.deref sign_ref) sorts;  | 
| 
2177
 
8b365a3a6ed1
Changed some mem, ins and union calls to be monomorphic
 
paulson 
parents: 
2163 
diff
changeset
 | 
533  | 
val shyps' = filter (fn S => mem_sort(S,sorts) orelse maybe_empty S) shyps;  | 
| 1238 | 534  | 
in  | 
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
535  | 
    Thm {sign_ref = sign_ref, der = der, maxidx = maxidx,
 | 
| 2386 | 536  | 
shyps =  | 
537  | 
(if eq_set_sort (shyps',sorts) orelse  | 
|
538  | 
not (!force_strip_shyps) then shyps'  | 
|
| 3061 | 539  | 
else (* FIXME tmp (since 1995/08/01) *)  | 
| 2386 | 540  | 
              (warning ("Removed sort hypotheses: " ^
 | 
| 2962 | 541  | 
commas (map Sorts.str_of_sort (shyps' \\ sorts)));  | 
| 2386 | 542  | 
warning "Let's hope these sorts are non-empty!";  | 
| 1238 | 543  | 
sorts)),  | 
| 1529 | 544  | 
hyps = hyps,  | 
545  | 
prop = prop}  | 
|
| 1238 | 546  | 
end;  | 
547  | 
||
548  | 
||
549  | 
(* implies_intr_shyps *)  | 
|
550  | 
||
551  | 
(*discharge all extra sort hypotheses*)  | 
|
552  | 
fun implies_intr_shyps thm =  | 
|
553  | 
(case extra_shyps thm of  | 
|
554  | 
[] => thm  | 
|
555  | 
| xshyps =>  | 
|
556  | 
let  | 
|
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
557  | 
        val Thm {sign_ref, der, maxidx, shyps, hyps, prop} = thm;
 | 
| 
2182
 
29e56f003599
Removal of polymorphic equality via mem, subset, eq_set, etc
 
paulson 
parents: 
2177 
diff
changeset
 | 
558  | 
val shyps' = ins_sort (logicS, shyps \\ xshyps);  | 
| 1238 | 559  | 
val used_names = foldr add_term_tfree_names (prop :: hyps, []);  | 
560  | 
val names =  | 
|
561  | 
tl (variantlist (replicate (length xshyps + 1) "'", used_names));  | 
|
562  | 
val tfrees = map (TFree o rpair logicS) names;  | 
|
563  | 
||
564  | 
fun mk_insort (T, S) = map (Logic.mk_inclass o pair T) S;  | 
|
| 2671 | 565  | 
val sort_hyps = List.concat (map2 mk_insort (tfrees, xshyps));  | 
| 1238 | 566  | 
in  | 
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
567  | 
        Thm {sign_ref = sign_ref, 
 | 
| 2386 | 568  | 
der = infer_derivs (Implies_intr_shyps, [der]),  | 
569  | 
maxidx = maxidx,  | 
|
570  | 
shyps = shyps',  | 
|
571  | 
hyps = hyps,  | 
|
572  | 
prop = Logic.list_implies (sort_hyps, prop)}  | 
|
| 1238 | 573  | 
end);  | 
574  | 
||
575  | 
||
| 1529 | 576  | 
(** Axioms **)  | 
| 
387
 
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
 
wenzelm 
parents: 
309 
diff
changeset
 | 
577  | 
|
| 
 
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
 
wenzelm 
parents: 
309 
diff
changeset
 | 
578  | 
(*look up the named axiom in the theory*)  | 
| 3812 | 579  | 
fun get_axiom theory raw_name =  | 
| 
387
 
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
 
wenzelm 
parents: 
309 
diff
changeset
 | 
580  | 
let  | 
| 3994 | 581  | 
val name = Sign.intern (sign_of theory) Theory.axiomK raw_name;  | 
| 
387
 
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
 
wenzelm 
parents: 
309 
diff
changeset
 | 
582  | 
fun get_ax [] = raise Match  | 
| 1529 | 583  | 
| get_ax (thy :: thys) =  | 
| 3994 | 584  | 
          let val {sign, axioms, parents, ...} = rep_theory thy
 | 
585  | 
in case Symtab.lookup (axioms, name) of  | 
|
| 2386 | 586  | 
Some t => fix_shyps [] []  | 
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
587  | 
                           (Thm {sign_ref = Sign.self_ref sign,
 | 
| 2386 | 588  | 
der = infer_derivs (Axiom(theory,name), []),  | 
589  | 
maxidx = maxidx_of_term t,  | 
|
590  | 
shyps = [],  | 
|
591  | 
hyps = [],  | 
|
592  | 
prop = t})  | 
|
593  | 
| None => get_ax parents handle Match => get_ax thys  | 
|
| 1529 | 594  | 
end;  | 
| 
387
 
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
 
wenzelm 
parents: 
309 
diff
changeset
 | 
595  | 
in  | 
| 
 
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
 
wenzelm 
parents: 
309 
diff
changeset
 | 
596  | 
get_ax [theory] handle Match  | 
| 3994 | 597  | 
      => raise THEORY ("No axiom " ^ quote name, [theory])
 | 
| 
387
 
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
 
wenzelm 
parents: 
309 
diff
changeset
 | 
598  | 
end;  | 
| 
 
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
 
wenzelm 
parents: 
309 
diff
changeset
 | 
599  | 
|
| 1529 | 600  | 
|
| 
776
 
df8f91c0e57c
improved axioms_of: returns thms as the manual says;
 
wenzelm 
parents: 
721 
diff
changeset
 | 
601  | 
(*return additional axioms of this theory node*)  | 
| 
 
df8f91c0e57c
improved axioms_of: returns thms as the manual says;
 
wenzelm 
parents: 
721 
diff
changeset
 | 
602  | 
fun axioms_of thy =  | 
| 
 
df8f91c0e57c
improved axioms_of: returns thms as the manual says;
 
wenzelm 
parents: 
721 
diff
changeset
 | 
603  | 
map (fn (s, _) => (s, get_axiom thy s))  | 
| 3994 | 604  | 
(Symtab.dest (#axioms (rep_theory thy)));  | 
| 
776
 
df8f91c0e57c
improved axioms_of: returns thms as the manual says;
 
wenzelm 
parents: 
721 
diff
changeset
 | 
605  | 
|
| 
1597
 
54ece585bf62
name_thm no longer takes a theory argument, as the
 
paulson 
parents: 
1580 
diff
changeset
 | 
606  | 
(*Attach a label to a theorem to make proof objects more readable*)  | 
| 4018 | 607  | 
fun name_thm (name, th as Thm {sign_ref, der, maxidx, shyps, hyps, prop}) =
 | 
608  | 
(case der of  | 
|
609  | 
Join (Theorem _, _) => th  | 
|
610  | 
| Join (Axiom _, _) => th  | 
|
611  | 
  | _ => Thm {sign_ref = sign_ref, der = Join (Theorem name, [der]),
 | 
|
612  | 
maxidx = maxidx, shyps = shyps, hyps = hyps, prop = prop});  | 
|
613  | 
||
614  | 
fun name_of_thm (Thm {der, ...}) =
 | 
|
615  | 
(case der of  | 
|
616  | 
Join (Theorem name, _) => name  | 
|
617  | 
| Join (Axiom (_, name), _) => name  | 
|
618  | 
| _ => "");  | 
|
| 0 | 619  | 
|
620  | 
||
| 1529 | 621  | 
(*Compression of theorems -- a separate rule, not integrated with the others,  | 
622  | 
as it could be slow.*)  | 
|
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
623  | 
fun compress (Thm {sign_ref, der, maxidx, shyps, hyps, prop}) = 
 | 
| 
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
624  | 
    Thm {sign_ref = sign_ref, 
 | 
| 2386 | 625  | 
der = der, (*No derivation recorded!*)  | 
626  | 
maxidx = maxidx,  | 
|
627  | 
shyps = shyps,  | 
|
628  | 
hyps = map Term.compress_term hyps,  | 
|
629  | 
prop = Term.compress_term prop};  | 
|
| 564 | 630  | 
|
| 
387
 
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
 
wenzelm 
parents: 
309 
diff
changeset
 | 
631  | 
|
| 2509 | 632  | 
|
| 1529 | 633  | 
(*** Meta rules ***)  | 
| 0 | 634  | 
|
| 2147 | 635  | 
(*Check that term does not contain same var with different typing/sorting.  | 
636  | 
If this check must be made, recalculate maxidx in hope of preventing its  | 
|
637  | 
recurrence.*)  | 
|
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
638  | 
fun nodup_Vars (thm as Thm{sign_ref, der, maxidx, shyps, hyps, prop}) s =
 | 
| 2147 | 639  | 
(Sign.nodup_Vars prop;  | 
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
640  | 
   Thm {sign_ref = sign_ref, 
 | 
| 2386 | 641  | 
der = der,  | 
642  | 
maxidx = maxidx_of_term prop,  | 
|
643  | 
shyps = shyps,  | 
|
644  | 
hyps = hyps,  | 
|
645  | 
prop = prop})  | 
|
| 2147 | 646  | 
handle TYPE(msg,Ts,ts) => raise TYPE(s^": "^msg,Ts,ts);  | 
| 
1495
 
b8b54847c77f
Added check for duplicate vars with distinct types/sorts (nodup_Vars)
 
nipkow 
parents: 
1493 
diff
changeset
 | 
647  | 
|
| 1220 | 648  | 
(** 'primitive' rules **)  | 
649  | 
||
650  | 
(*discharge all assumptions t from ts*)  | 
|
| 0 | 651  | 
val disch = gen_rem (op aconv);  | 
652  | 
||
| 1220 | 653  | 
(*The assumption rule A|-A in a theory*)  | 
| 250 | 654  | 
fun assume ct : thm =  | 
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
655  | 
  let val Cterm {sign_ref, t=prop, T, maxidx} = ct
 | 
| 250 | 656  | 
in if T<>propT then  | 
657  | 
        raise THM("assume: assumptions must have type prop", 0, [])
 | 
|
| 0 | 658  | 
else if maxidx <> ~1 then  | 
| 250 | 659  | 
        raise THM("assume: assumptions may not contain scheme variables",
 | 
660  | 
maxidx, [])  | 
|
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
661  | 
      else Thm{sign_ref   = sign_ref,
 | 
| 2386 | 662  | 
der = infer_derivs (Assume ct, []),  | 
663  | 
maxidx = ~1,  | 
|
664  | 
shyps = add_term_sorts(prop,[]),  | 
|
665  | 
hyps = [prop],  | 
|
666  | 
prop = prop}  | 
|
| 0 | 667  | 
end;  | 
668  | 
||
| 1220 | 669  | 
(*Implication introduction  | 
| 3529 | 670  | 
[A]  | 
671  | 
:  | 
|
672  | 
B  | 
|
| 1220 | 673  | 
-------  | 
674  | 
A ==> B  | 
|
675  | 
*)  | 
|
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
676  | 
fun implies_intr cA (thB as Thm{sign_ref,der,maxidx,hyps,prop,...}) : thm =
 | 
| 
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
677  | 
  let val Cterm {sign_ref=sign_refA, t=A, T, maxidx=maxidxA} = cA
 | 
| 0 | 678  | 
in if T<>propT then  | 
| 250 | 679  | 
        raise THM("implies_intr: assumptions must have type prop", 0, [thB])
 | 
| 1238 | 680  | 
else fix_shyps [thB] []  | 
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
681  | 
        (Thm{sign_ref = Sign.merge_refs (sign_ref,sign_refA),  
 | 
| 2386 | 682  | 
der = infer_derivs (Implies_intr cA, [der]),  | 
683  | 
maxidx = Int.max(maxidxA, maxidx),  | 
|
684  | 
shyps = [],  | 
|
685  | 
hyps = disch(hyps,A),  | 
|
686  | 
prop = implies$A$prop})  | 
|
| 0 | 687  | 
handle TERM _ =>  | 
688  | 
        raise THM("implies_intr: incompatible signatures", 0, [thB])
 | 
|
689  | 
end;  | 
|
690  | 
||
| 1529 | 691  | 
|
| 1220 | 692  | 
(*Implication elimination  | 
693  | 
A ==> B A  | 
|
694  | 
------------  | 
|
695  | 
B  | 
|
696  | 
*)  | 
|
| 0 | 697  | 
fun implies_elim thAB thA : thm =  | 
| 1529 | 698  | 
    let val Thm{maxidx=maxA, der=derA, hyps=hypsA, prop=propA,...} = thA
 | 
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
699  | 
        and Thm{sign_ref, der, maxidx, hyps, prop,...} = thAB;
 | 
| 250 | 700  | 
        fun err(a) = raise THM("implies_elim: "^a, 0, [thAB,thA])
 | 
| 0 | 701  | 
in case prop of  | 
| 250 | 702  | 
imp$A$B =>  | 
703  | 
if imp=implies andalso A aconv propA  | 
|
| 1220 | 704  | 
then fix_shyps [thAB, thA] []  | 
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
705  | 
                       (Thm{sign_ref= merge_thm_sgs(thAB,thA),
 | 
| 2386 | 706  | 
der = infer_derivs (Implies_elim, [der,derA]),  | 
707  | 
maxidx = Int.max(maxA,maxidx),  | 
|
708  | 
shyps = [],  | 
|
709  | 
hyps = union_term(hypsA,hyps), (*dups suppressed*)  | 
|
710  | 
prop = B})  | 
|
| 250 | 711  | 
                else err("major premise")
 | 
712  | 
          | _ => err("major premise")
 | 
|
| 0 | 713  | 
end;  | 
| 250 | 714  | 
|
| 1220 | 715  | 
(*Forall introduction. The Free or Var x must not be free in the hypotheses.  | 
716  | 
A  | 
|
717  | 
-----  | 
|
718  | 
!!x.A  | 
|
719  | 
*)  | 
|
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
720  | 
fun forall_intr cx (th as Thm{sign_ref,der,maxidx,hyps,prop,...}) =
 | 
| 
229
 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 
lcp 
parents: 
225 
diff
changeset
 | 
721  | 
let val x = term_of cx;  | 
| 1238 | 722  | 
fun result(a,T) = fix_shyps [th] []  | 
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
723  | 
        (Thm{sign_ref = sign_ref, 
 | 
| 2386 | 724  | 
der = infer_derivs (Forall_intr cx, [der]),  | 
725  | 
maxidx = maxidx,  | 
|
726  | 
shyps = [],  | 
|
727  | 
hyps = hyps,  | 
|
728  | 
prop = all(T) $ Abs(a, T, abstract_over (x,prop))})  | 
|
| 0 | 729  | 
in case x of  | 
| 250 | 730  | 
Free(a,T) =>  | 
731  | 
if exists (apl(x, Logic.occs)) hyps  | 
|
732  | 
          then  raise THM("forall_intr: variable free in assumptions", 0, [th])
 | 
|
733  | 
else result(a,T)  | 
|
| 0 | 734  | 
| Var((a,_),T) => result(a,T)  | 
735  | 
      | _ => raise THM("forall_intr: not a variable", 0, [th])
 | 
|
736  | 
end;  | 
|
737  | 
||
| 1220 | 738  | 
(*Forall elimination  | 
739  | 
!!x.A  | 
|
740  | 
------  | 
|
741  | 
A[t/x]  | 
|
742  | 
*)  | 
|
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
743  | 
fun forall_elim ct (th as Thm{sign_ref,der,maxidx,hyps,prop,...}) : thm =
 | 
| 
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
744  | 
  let val Cterm {sign_ref=sign_reft, t, T, maxidx=maxt} = ct
 | 
| 0 | 745  | 
in case prop of  | 
| 2386 | 746  | 
        Const("all",Type("fun",[Type("fun",[qary,_]),_])) $ A =>
 | 
747  | 
if T<>qary then  | 
|
748  | 
              raise THM("forall_elim: type mismatch", 0, [th])
 | 
|
749  | 
else let val thm = fix_shyps [th] []  | 
|
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
750  | 
                    (Thm{sign_ref= Sign.merge_refs(sign_ref,sign_reft),
 | 
| 2386 | 751  | 
der = infer_derivs (Forall_elim ct, [der]),  | 
752  | 
maxidx = Int.max(maxidx, maxt),  | 
|
753  | 
shyps = [],  | 
|
754  | 
hyps = hyps,  | 
|
755  | 
prop = betapply(A,t)})  | 
|
756  | 
in if maxt >= 0 andalso maxidx >= 0  | 
|
757  | 
then nodup_Vars thm "forall_elim"  | 
|
758  | 
else thm (*no new Vars: no expensive check!*)  | 
|
759  | 
end  | 
|
| 2147 | 760  | 
      | _ => raise THM("forall_elim: not quantified", 0, [th])
 | 
| 0 | 761  | 
end  | 
762  | 
handle TERM _ =>  | 
|
| 250 | 763  | 
         raise THM("forall_elim: incompatible signatures", 0, [th]);
 | 
| 0 | 764  | 
|
765  | 
||
| 1220 | 766  | 
(* Equality *)  | 
| 0 | 767  | 
|
768  | 
(*The reflexivity rule: maps t to the theorem t==t *)  | 
|
| 250 | 769  | 
fun reflexive ct =  | 
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
770  | 
  let val Cterm {sign_ref, t, T, maxidx} = ct
 | 
| 1238 | 771  | 
in fix_shyps [] []  | 
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
772  | 
       (Thm{sign_ref= sign_ref, 
 | 
| 2386 | 773  | 
der = infer_derivs (Reflexive ct, []),  | 
774  | 
shyps = [],  | 
|
775  | 
hyps = [],  | 
|
776  | 
maxidx = maxidx,  | 
|
777  | 
prop = Logic.mk_equals(t,t)})  | 
|
| 0 | 778  | 
end;  | 
779  | 
||
780  | 
(*The symmetry rule  | 
|
| 1220 | 781  | 
t==u  | 
782  | 
----  | 
|
783  | 
u==t  | 
|
784  | 
*)  | 
|
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
785  | 
fun symmetric (th as Thm{sign_ref,der,maxidx,shyps,hyps,prop}) =
 | 
| 0 | 786  | 
case prop of  | 
787  | 
      (eq as Const("==",_)) $ t $ u =>
 | 
|
| 1238 | 788  | 
(*no fix_shyps*)  | 
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
789  | 
          Thm{sign_ref = sign_ref,
 | 
| 2386 | 790  | 
der = infer_derivs (Symmetric, [der]),  | 
791  | 
maxidx = maxidx,  | 
|
792  | 
shyps = shyps,  | 
|
793  | 
hyps = hyps,  | 
|
794  | 
prop = eq$u$t}  | 
|
| 0 | 795  | 
    | _ => raise THM("symmetric", 0, [th]);
 | 
796  | 
||
797  | 
(*The transitive rule  | 
|
| 1220 | 798  | 
t1==u u==t2  | 
799  | 
--------------  | 
|
800  | 
t1==t2  | 
|
801  | 
*)  | 
|
| 0 | 802  | 
fun transitive th1 th2 =  | 
| 1529 | 803  | 
  let val Thm{der=der1, maxidx=max1, hyps=hyps1, prop=prop1,...} = th1
 | 
804  | 
      and Thm{der=der2, maxidx=max2, hyps=hyps2, prop=prop2,...} = th2;
 | 
|
| 0 | 805  | 
      fun err(msg) = raise THM("transitive: "^msg, 0, [th1,th2])
 | 
806  | 
in case (prop1,prop2) of  | 
|
807  | 
       ((eq as Const("==",_)) $ t1 $ u, Const("==",_) $ u' $ t2) =>
 | 
|
| 1634 | 808  | 
if not (u aconv u') then err"middle term"  | 
809  | 
else let val thm =  | 
|
| 1220 | 810  | 
fix_shyps [th1, th2] []  | 
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
811  | 
                (Thm{sign_ref= merge_thm_sgs(th1,th2), 
 | 
| 2386 | 812  | 
der = infer_derivs (Transitive, [der1, der2]),  | 
| 2147 | 813  | 
maxidx = Int.max(max1,max2),  | 
| 2386 | 814  | 
shyps = [],  | 
815  | 
hyps = union_term(hyps1,hyps2),  | 
|
816  | 
prop = eq$t1$t2})  | 
|
| 
2139
 
2c59b204b540
Only calls nodup_Vars if really necessary.  We get a speedup of nearly 6%
 
paulson 
parents: 
2047 
diff
changeset
 | 
817  | 
in if max1 >= 0 andalso max2 >= 0  | 
| 2147 | 818  | 
then nodup_Vars thm "transitive"  | 
819  | 
else thm (*no new Vars: no expensive check!*)  | 
|
| 
2139
 
2c59b204b540
Only calls nodup_Vars if really necessary.  We get a speedup of nearly 6%
 
paulson 
parents: 
2047 
diff
changeset
 | 
820  | 
end  | 
| 0 | 821  | 
| _ => err"premises"  | 
822  | 
end;  | 
|
823  | 
||
| 1160 | 824  | 
(*Beta-conversion: maps (%x.t)(u) to the theorem (%x.t)(u) == t[u/x] *)  | 
| 250 | 825  | 
fun beta_conversion ct =  | 
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
826  | 
  let val Cterm {sign_ref, t, T, maxidx} = ct
 | 
| 0 | 827  | 
in case t of  | 
| 1238 | 828  | 
Abs(_,_,bodt) $ u => fix_shyps [] []  | 
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
829  | 
            (Thm{sign_ref = sign_ref,  
 | 
| 2386 | 830  | 
der = infer_derivs (Beta_conversion ct, []),  | 
831  | 
maxidx = maxidx,  | 
|
832  | 
shyps = [],  | 
|
833  | 
hyps = [],  | 
|
834  | 
prop = Logic.mk_equals(t, subst_bound (u,bodt))})  | 
|
| 250 | 835  | 
        | _ =>  raise THM("beta_conversion: not a redex", 0, [])
 | 
| 0 | 836  | 
end;  | 
837  | 
||
838  | 
(*The extensionality rule (proviso: x not free in f, g, or hypotheses)  | 
|
| 1220 | 839  | 
f(x) == g(x)  | 
840  | 
------------  | 
|
841  | 
f == g  | 
|
842  | 
*)  | 
|
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
843  | 
fun extensional (th as Thm{sign_ref, der, maxidx,shyps,hyps,prop}) =
 | 
| 0 | 844  | 
case prop of  | 
845  | 
    (Const("==",_)) $ (f$x) $ (g$y) =>
 | 
|
| 250 | 846  | 
      let fun err(msg) = raise THM("extensional: "^msg, 0, [th])
 | 
| 0 | 847  | 
in (if x<>y then err"different variables" else  | 
848  | 
case y of  | 
|
| 250 | 849  | 
Free _ =>  | 
850  | 
if exists (apl(y, Logic.occs)) (f::g::hyps)  | 
|
851  | 
then err"variable free in hyps or functions" else ()  | 
|
852  | 
| Var _ =>  | 
|
853  | 
if Logic.occs(y,f) orelse Logic.occs(y,g)  | 
|
854  | 
then err"variable free in functions" else ()  | 
|
855  | 
| _ => err"not a variable");  | 
|
| 1238 | 856  | 
(*no fix_shyps*)  | 
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
857  | 
          Thm{sign_ref = sign_ref,
 | 
| 2386 | 858  | 
der = infer_derivs (Extensional, [der]),  | 
859  | 
maxidx = maxidx,  | 
|
860  | 
shyps = shyps,  | 
|
861  | 
hyps = hyps,  | 
|
| 1529 | 862  | 
prop = Logic.mk_equals(f,g)}  | 
| 0 | 863  | 
end  | 
864  | 
 | _ =>  raise THM("extensional: premise", 0, [th]);
 | 
|
865  | 
||
866  | 
(*The abstraction rule. The Free or Var x must not be free in the hypotheses.  | 
|
867  | 
The bound variable will be named "a" (since x will be something like x320)  | 
|
| 1220 | 868  | 
t == u  | 
869  | 
------------  | 
|
870  | 
%x.t == %x.u  | 
|
871  | 
*)  | 
|
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
872  | 
fun abstract_rule a cx (th as Thm{sign_ref,der,maxidx,hyps,prop,...}) =
 | 
| 
229
 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 
lcp 
parents: 
225 
diff
changeset
 | 
873  | 
let val x = term_of cx;  | 
| 250 | 874  | 
val (t,u) = Logic.dest_equals prop  | 
875  | 
handle TERM _ =>  | 
|
876  | 
                raise THM("abstract_rule: premise not an equality", 0, [th])
 | 
|
| 1238 | 877  | 
fun result T = fix_shyps [th] []  | 
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
878  | 
          (Thm{sign_ref = sign_ref,
 | 
| 2386 | 879  | 
der = infer_derivs (Abstract_rule (a,cx), [der]),  | 
880  | 
maxidx = maxidx,  | 
|
881  | 
shyps = [],  | 
|
882  | 
hyps = hyps,  | 
|
883  | 
prop = Logic.mk_equals(Abs(a, T, abstract_over (x,t)),  | 
|
884  | 
Abs(a, T, abstract_over (x,u)))})  | 
|
| 0 | 885  | 
in case x of  | 
| 250 | 886  | 
Free(_,T) =>  | 
887  | 
if exists (apl(x, Logic.occs)) hyps  | 
|
888  | 
         then raise THM("abstract_rule: variable free in assumptions", 0, [th])
 | 
|
889  | 
else result T  | 
|
| 0 | 890  | 
| Var(_,T) => result T  | 
891  | 
      | _ => raise THM("abstract_rule: not a variable", 0, [th])
 | 
|
892  | 
end;  | 
|
893  | 
||
894  | 
(*The combination rule  | 
|
| 3529 | 895  | 
f == g t == u  | 
896  | 
--------------  | 
|
897  | 
f(t) == g(u)  | 
|
| 1220 | 898  | 
*)  | 
| 0 | 899  | 
fun combination th1 th2 =  | 
| 1529 | 900  | 
  let val Thm{der=der1, maxidx=max1, shyps=shyps1, hyps=hyps1, 
 | 
| 2386 | 901  | 
prop=prop1,...} = th1  | 
| 1529 | 902  | 
      and Thm{der=der2, maxidx=max2, shyps=shyps2, hyps=hyps2, 
 | 
| 2386 | 903  | 
prop=prop2,...} = th2  | 
| 
1836
 
861e29c7cada
Added type-checking to rule "combination".  This corrects a fault
 
paulson 
parents: 
1802 
diff
changeset
 | 
904  | 
fun chktypes (f,t) =  | 
| 2386 | 905  | 
(case fastype_of f of  | 
906  | 
                Type("fun",[T1,T2]) => 
 | 
|
907  | 
if T1 <> fastype_of t then  | 
|
908  | 
                         raise THM("combination: types", 0, [th1,th2])
 | 
|
909  | 
else ()  | 
|
910  | 
                | _ => raise THM("combination: not function type", 0, 
 | 
|
911  | 
[th1,th2]))  | 
|
| 
1495
 
b8b54847c77f
Added check for duplicate vars with distinct types/sorts (nodup_Vars)
 
nipkow 
parents: 
1493 
diff
changeset
 | 
912  | 
in case (prop1,prop2) of  | 
| 0 | 913  | 
       (Const("==",_) $ f $ g, Const("==",_) $ t $ u) =>
 | 
| 
1836
 
861e29c7cada
Added type-checking to rule "combination".  This corrects a fault
 
paulson 
parents: 
1802 
diff
changeset
 | 
914  | 
let val _ = chktypes (f,t)  | 
| 2386 | 915  | 
val thm = (*no fix_shyps*)  | 
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
916  | 
                        Thm{sign_ref = merge_thm_sgs(th1,th2), 
 | 
| 2386 | 917  | 
der = infer_derivs (Combination, [der1, der2]),  | 
918  | 
maxidx = Int.max(max1,max2),  | 
|
919  | 
shyps = union_sort(shyps1,shyps2),  | 
|
920  | 
hyps = union_term(hyps1,hyps2),  | 
|
921  | 
prop = Logic.mk_equals(f$t, g$u)}  | 
|
| 
2139
 
2c59b204b540
Only calls nodup_Vars if really necessary.  We get a speedup of nearly 6%
 
paulson 
parents: 
2047 
diff
changeset
 | 
922  | 
in if max1 >= 0 andalso max2 >= 0  | 
| 
 
2c59b204b540
Only calls nodup_Vars if really necessary.  We get a speedup of nearly 6%
 
paulson 
parents: 
2047 
diff
changeset
 | 
923  | 
then nodup_Vars thm "combination"  | 
| 2386 | 924  | 
else thm (*no new Vars: no expensive check!*)  | 
| 
2139
 
2c59b204b540
Only calls nodup_Vars if really necessary.  We get a speedup of nearly 6%
 
paulson 
parents: 
2047 
diff
changeset
 | 
925  | 
end  | 
| 0 | 926  | 
     | _ =>  raise THM("combination: premises", 0, [th1,th2])
 | 
927  | 
end;  | 
|
928  | 
||
929  | 
||
930  | 
(* Equality introduction  | 
|
| 3529 | 931  | 
A ==> B B ==> A  | 
932  | 
----------------  | 
|
933  | 
A == B  | 
|
| 1220 | 934  | 
*)  | 
| 0 | 935  | 
fun equal_intr th1 th2 =  | 
| 1529 | 936  | 
  let val Thm{der=der1,maxidx=max1, shyps=shyps1, hyps=hyps1, 
 | 
| 2386 | 937  | 
prop=prop1,...} = th1  | 
| 1529 | 938  | 
      and Thm{der=der2, maxidx=max2, shyps=shyps2, hyps=hyps2, 
 | 
| 2386 | 939  | 
prop=prop2,...} = th2;  | 
| 1529 | 940  | 
      fun err(msg) = raise THM("equal_intr: "^msg, 0, [th1,th2])
 | 
941  | 
in case (prop1,prop2) of  | 
|
942  | 
       (Const("==>",_) $ A $ B, Const("==>",_) $ B' $ A')  =>
 | 
|
| 2386 | 943  | 
if A aconv A' andalso B aconv B'  | 
944  | 
then  | 
|
945  | 
(*no fix_shyps*)  | 
|
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
946  | 
              Thm{sign_ref = merge_thm_sgs(th1,th2),
 | 
| 2386 | 947  | 
der = infer_derivs (Equal_intr, [der1, der2]),  | 
948  | 
maxidx = Int.max(max1,max2),  | 
|
949  | 
shyps = union_sort(shyps1,shyps2),  | 
|
950  | 
hyps = union_term(hyps1,hyps2),  | 
|
951  | 
prop = Logic.mk_equals(A,B)}  | 
|
952  | 
else err"not equal"  | 
|
| 1529 | 953  | 
| _ => err"premises"  | 
954  | 
end;  | 
|
955  | 
||
956  | 
||
957  | 
(*The equal propositions rule  | 
|
| 3529 | 958  | 
A == B A  | 
| 1529 | 959  | 
---------  | 
960  | 
B  | 
|
961  | 
*)  | 
|
962  | 
fun equal_elim th1 th2 =  | 
|
963  | 
  let val Thm{der=der1, maxidx=max1, hyps=hyps1, prop=prop1,...} = th1
 | 
|
964  | 
      and Thm{der=der2, maxidx=max2, hyps=hyps2, prop=prop2,...} = th2;
 | 
|
965  | 
      fun err(msg) = raise THM("equal_elim: "^msg, 0, [th1,th2])
 | 
|
966  | 
in case prop1 of  | 
|
967  | 
       Const("==",_) $ A $ B =>
 | 
|
968  | 
if not (prop2 aconv A) then err"not equal" else  | 
|
969  | 
fix_shyps [th1, th2] []  | 
|
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
970  | 
              (Thm{sign_ref= merge_thm_sgs(th1,th2), 
 | 
| 2386 | 971  | 
der = infer_derivs (Equal_elim, [der1, der2]),  | 
972  | 
maxidx = Int.max(max1,max2),  | 
|
973  | 
shyps = [],  | 
|
974  | 
hyps = union_term(hyps1,hyps2),  | 
|
975  | 
prop = B})  | 
|
| 1529 | 976  | 
| _ => err"major premise"  | 
977  | 
end;  | 
|
| 0 | 978  | 
|
| 1220 | 979  | 
|
980  | 
||
| 0 | 981  | 
(**** Derived rules ****)  | 
982  | 
||
| 1503 | 983  | 
(*Discharge all hypotheses. Need not verify cterms or call fix_shyps.  | 
| 0 | 984  | 
Repeated hypotheses are discharged only once; fold cannot do this*)  | 
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
985  | 
fun implies_intr_hyps (Thm{sign_ref, der, maxidx, shyps, hyps=A::As, prop}) =
 | 
| 1238 | 986  | 
implies_intr_hyps (*no fix_shyps*)  | 
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
987  | 
            (Thm{sign_ref = sign_ref, 
 | 
| 2386 | 988  | 
der = infer_derivs (Implies_intr_hyps, [der]),  | 
989  | 
maxidx = maxidx,  | 
|
990  | 
shyps = shyps,  | 
|
| 1529 | 991  | 
hyps = disch(As,A),  | 
| 2386 | 992  | 
prop = implies$A$prop})  | 
| 0 | 993  | 
| implies_intr_hyps th = th;  | 
994  | 
||
995  | 
(*Smash" unifies the list of term pairs leaving no flex-flex pairs.  | 
|
| 250 | 996  | 
Instantiates the theorem and deletes trivial tpairs.  | 
| 0 | 997  | 
Resulting sequence may contain multiple elements if the tpairs are  | 
998  | 
not all flex-flex. *)  | 
|
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
999  | 
fun flexflex_rule (th as Thm{sign_ref, der, maxidx, hyps, prop,...}) =
 | 
| 250 | 1000  | 
let fun newthm env =  | 
| 1529 | 1001  | 
if Envir.is_empty env then th  | 
1002  | 
else  | 
|
| 250 | 1003  | 
let val (tpairs,horn) =  | 
1004  | 
Logic.strip_flexpairs (Envir.norm_term env prop)  | 
|
1005  | 
(*Remove trivial tpairs, of the form t=t*)  | 
|
1006  | 
val distpairs = filter (not o op aconv) tpairs  | 
|
1007  | 
val newprop = Logic.list_flexpairs(distpairs, horn)  | 
|
| 1220 | 1008  | 
in fix_shyps [th] (env_codT env)  | 
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1009  | 
                (Thm{sign_ref = sign_ref, 
 | 
| 2386 | 1010  | 
der = infer_derivs (Flexflex_rule env, [der]),  | 
1011  | 
maxidx = maxidx_of_term newprop,  | 
|
1012  | 
shyps = [],  | 
|
1013  | 
hyps = hyps,  | 
|
1014  | 
prop = newprop})  | 
|
| 250 | 1015  | 
end;  | 
| 0 | 1016  | 
val (tpairs,_) = Logic.strip_flexpairs prop  | 
1017  | 
in Sequence.maps newthm  | 
|
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1018  | 
(Unify.smash_unifiers(Sign.deref sign_ref, Envir.empty maxidx, tpairs))  | 
| 0 | 1019  | 
end;  | 
1020  | 
||
1021  | 
(*Instantiation of Vars  | 
|
| 1220 | 1022  | 
A  | 
1023  | 
-------------------  | 
|
1024  | 
A[t1/v1,....,tn/vn]  | 
|
1025  | 
*)  | 
|
| 0 | 1026  | 
|
1027  | 
(*Check that all the terms are Vars and are distinct*)  | 
|
1028  | 
fun instl_ok ts = forall is_Var ts andalso null(findrep ts);  | 
|
1029  | 
||
1030  | 
(*For instantiate: process pair of cterms, merge theories*)  | 
|
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1031  | 
fun add_ctpair ((ct,cu), (sign_ref,tpairs)) =  | 
| 
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1032  | 
  let val Cterm {sign_ref=sign_reft, t=t, T= T, ...} = ct
 | 
| 
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1033  | 
      and Cterm {sign_ref=sign_refu, t=u, T= U, ...} = cu
 | 
| 
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1034  | 
in  | 
| 
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1035  | 
if T=U then  | 
| 
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1036  | 
(Sign.merge_refs (sign_ref, Sign.merge_refs (sign_reft, sign_refu)), (t,u)::tpairs)  | 
| 
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1037  | 
    else raise TYPE("add_ctpair", [T,U], [t,u])
 | 
| 0 | 1038  | 
end;  | 
1039  | 
||
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1040  | 
fun add_ctyp ((v,ctyp), (sign_ref',vTs)) =  | 
| 
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1041  | 
  let val Ctyp {T,sign_ref} = ctyp
 | 
| 
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1042  | 
in (Sign.merge_refs(sign_ref,sign_ref'), (v,T)::vTs) end;  | 
| 0 | 1043  | 
|
1044  | 
(*Left-to-right replacements: ctpairs = [...,(vi,ti),...].  | 
|
1045  | 
Instantiates distinct Vars by terms of same type.  | 
|
1046  | 
Normalizes the new theorem! *)  | 
|
| 1529 | 1047  | 
fun instantiate ([], []) th = th  | 
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1048  | 
  | instantiate (vcTs,ctpairs)  (th as Thm{sign_ref,der,maxidx,hyps,prop,...}) =
 | 
| 
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1049  | 
let val (newsign_ref,tpairs) = foldr add_ctpair (ctpairs, (sign_ref,[]));  | 
| 
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1050  | 
val (newsign_ref,vTs) = foldr add_ctyp (vcTs, (newsign_ref,[]));  | 
| 250 | 1051  | 
val newprop =  | 
1052  | 
Envir.norm_term (Envir.empty 0)  | 
|
1053  | 
(subst_atomic tpairs  | 
|
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1054  | 
(Type.inst_term_tvars(Sign.tsig_of (Sign.deref newsign_ref),vTs) prop))  | 
| 1220 | 1055  | 
val newth =  | 
1056  | 
fix_shyps [th] (map snd vTs)  | 
|
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1057  | 
              (Thm{sign_ref = newsign_ref, 
 | 
| 2386 | 1058  | 
der = infer_derivs (Instantiate(vcTs,ctpairs), [der]),  | 
1059  | 
maxidx = maxidx_of_term newprop,  | 
|
1060  | 
shyps = [],  | 
|
1061  | 
hyps = hyps,  | 
|
1062  | 
prop = newprop})  | 
|
| 250 | 1063  | 
in if not(instl_ok(map #1 tpairs))  | 
| 193 | 1064  | 
      then raise THM("instantiate: variables not distinct", 0, [th])
 | 
1065  | 
else if not(null(findrep(map #1 vTs)))  | 
|
1066  | 
      then raise THM("instantiate: type variables not distinct", 0, [th])
 | 
|
| 2147 | 1067  | 
else nodup_Vars newth "instantiate"  | 
| 0 | 1068  | 
end  | 
| 250 | 1069  | 
handle TERM _ =>  | 
| 0 | 1070  | 
           raise THM("instantiate: incompatible signatures",0,[th])
 | 
| 2671 | 1071  | 
       | TYPE (msg,_,_) => raise THM("instantiate: type conflict: " ^ msg, 
 | 
1072  | 
0, [th]);  | 
|
| 0 | 1073  | 
|
1074  | 
(*The trivial implication A==>A, justified by assume and forall rules.  | 
|
1075  | 
A can contain Vars, not so for assume! *)  | 
|
| 250 | 1076  | 
fun trivial ct : thm =  | 
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1077  | 
  let val Cterm {sign_ref, t=A, T, maxidx} = ct
 | 
| 250 | 1078  | 
in if T<>propT then  | 
1079  | 
            raise THM("trivial: the term must have type prop", 0, [])
 | 
|
| 1238 | 1080  | 
else fix_shyps [] []  | 
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1081  | 
        (Thm{sign_ref = sign_ref, 
 | 
| 2386 | 1082  | 
der = infer_derivs (Trivial ct, []),  | 
1083  | 
maxidx = maxidx,  | 
|
1084  | 
shyps = [],  | 
|
1085  | 
hyps = [],  | 
|
1086  | 
prop = implies$A$A})  | 
|
| 0 | 1087  | 
end;  | 
1088  | 
||
| 1503 | 1089  | 
(*Axiom-scheme reflecting signature contents: "OFCLASS(?'a::c, c_class)" *)  | 
| 
399
 
86cc2b98f9e0
added class_triv: theory -> class -> thm (for axclasses);
 
wenzelm 
parents: 
387 
diff
changeset
 | 
1090  | 
fun class_triv thy c =  | 
| 1529 | 1091  | 
let val sign = sign_of thy;  | 
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1092  | 
      val Cterm {sign_ref, t, maxidx, ...} =
 | 
| 2386 | 1093  | 
          cterm_of sign (Logic.mk_inclass (TVar (("'a", 0), [c]), c))
 | 
1094  | 
            handle TERM (msg, _) => raise THM ("class_triv: " ^ msg, 0, []);
 | 
|
| 
399
 
86cc2b98f9e0
added class_triv: theory -> class -> thm (for axclasses);
 
wenzelm 
parents: 
387 
diff
changeset
 | 
1095  | 
in  | 
| 1238 | 1096  | 
fix_shyps [] []  | 
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1097  | 
      (Thm {sign_ref = sign_ref, 
 | 
| 2386 | 1098  | 
der = infer_derivs (Class_triv(thy,c), []),  | 
1099  | 
maxidx = maxidx,  | 
|
1100  | 
shyps = [],  | 
|
1101  | 
hyps = [],  | 
|
1102  | 
prop = t})  | 
|
| 
399
 
86cc2b98f9e0
added class_triv: theory -> class -> thm (for axclasses);
 
wenzelm 
parents: 
387 
diff
changeset
 | 
1103  | 
end;  | 
| 
 
86cc2b98f9e0
added class_triv: theory -> class -> thm (for axclasses);
 
wenzelm 
parents: 
387 
diff
changeset
 | 
1104  | 
|
| 
 
86cc2b98f9e0
added class_triv: theory -> class -> thm (for axclasses);
 
wenzelm 
parents: 
387 
diff
changeset
 | 
1105  | 
|
| 0 | 1106  | 
(* Replace all TFrees not in the hyps by new TVars *)  | 
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1107  | 
fun varifyT(Thm{sign_ref,der,maxidx,shyps,hyps,prop}) =
 | 
| 0 | 1108  | 
let val tfrees = foldr add_term_tfree_names (hyps,[])  | 
| 1634 | 1109  | 
in let val thm = (*no fix_shyps*)  | 
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1110  | 
    Thm{sign_ref = sign_ref, 
 | 
| 2386 | 1111  | 
der = infer_derivs (VarifyT, [der]),  | 
1112  | 
maxidx = Int.max(0,maxidx),  | 
|
1113  | 
shyps = shyps,  | 
|
1114  | 
hyps = hyps,  | 
|
| 1529 | 1115  | 
prop = Type.varify(prop,tfrees)}  | 
| 2147 | 1116  | 
in nodup_Vars thm "varifyT" end  | 
| 1634 | 1117  | 
(* this nodup_Vars check can be removed if thms are guaranteed not to contain  | 
1118  | 
duplicate TVars with differnt sorts *)  | 
|
| 0 | 1119  | 
end;  | 
1120  | 
||
1121  | 
(* Replace all TVars by new TFrees *)  | 
|
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1122  | 
fun freezeT(Thm{sign_ref,der,maxidx,shyps,hyps,prop}) =
 | 
| 3410 | 1123  | 
let val (prop',_) = Type.freeze_thaw prop  | 
| 1238 | 1124  | 
in (*no fix_shyps*)  | 
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1125  | 
    Thm{sign_ref = sign_ref, 
 | 
| 2386 | 1126  | 
der = infer_derivs (FreezeT, [der]),  | 
1127  | 
maxidx = maxidx_of_term prop',  | 
|
1128  | 
shyps = shyps,  | 
|
1129  | 
hyps = hyps,  | 
|
| 1529 | 1130  | 
prop = prop'}  | 
| 1220 | 1131  | 
end;  | 
| 0 | 1132  | 
|
1133  | 
||
1134  | 
(*** Inference rules for tactics ***)  | 
|
1135  | 
||
1136  | 
(*Destruct proof state into constraints, other goals, goal(i), rest *)  | 
|
1137  | 
fun dest_state (state as Thm{prop,...}, i) =
 | 
|
1138  | 
let val (tpairs,horn) = Logic.strip_flexpairs prop  | 
|
1139  | 
in case Logic.strip_prems(i, [], horn) of  | 
|
1140  | 
(B::rBs, C) => (tpairs, rev rBs, B, C)  | 
|
1141  | 
        | _ => raise THM("dest_state", i, [state])
 | 
|
1142  | 
end  | 
|
1143  | 
  handle TERM _ => raise THM("dest_state", i, [state]);
 | 
|
1144  | 
||
| 309 | 1145  | 
(*Increment variables and parameters of orule as required for  | 
| 0 | 1146  | 
resolution with goal i of state. *)  | 
1147  | 
fun lift_rule (state, i) orule =  | 
|
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1148  | 
  let val Thm{shyps=sshyps, prop=sprop, maxidx=smax, sign_ref=ssign_ref,...} = state
 | 
| 0 | 1149  | 
val (Bi::_, _) = Logic.strip_prems(i, [], Logic.skip_flexpairs sprop)  | 
| 1529 | 1150  | 
        handle TERM _ => raise THM("lift_rule", i, [orule,state])
 | 
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1151  | 
      val ct_Bi = Cterm {sign_ref=ssign_ref, maxidx=smax, T=propT, t=Bi}
 | 
| 1529 | 1152  | 
val (lift_abs,lift_all) = Logic.lift_fns(Bi,smax+1)  | 
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1153  | 
      val (Thm{sign_ref, der, maxidx,shyps,hyps,prop}) = orule
 | 
| 0 | 1154  | 
val (tpairs,As,B) = Logic.strip_horn prop  | 
| 1238 | 1155  | 
in (*no fix_shyps*)  | 
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1156  | 
      Thm{sign_ref = merge_thm_sgs(state,orule),
 | 
| 2386 | 1157  | 
der = infer_derivs (Lift_rule(ct_Bi, i), [der]),  | 
1158  | 
maxidx = maxidx+smax+1,  | 
|
| 
2177
 
8b365a3a6ed1
Changed some mem, ins and union calls to be monomorphic
 
paulson 
parents: 
2163 
diff
changeset
 | 
1159  | 
shyps=union_sort(sshyps,shyps),  | 
| 2386 | 1160  | 
hyps=hyps,  | 
| 1529 | 1161  | 
prop = Logic.rule_of (map (pairself lift_abs) tpairs,  | 
| 2386 | 1162  | 
map lift_all As,  | 
1163  | 
lift_all B)}  | 
|
| 0 | 1164  | 
end;  | 
1165  | 
||
1166  | 
(*Solve subgoal Bi of proof state B1...Bn/C by assumption. *)  | 
|
1167  | 
fun assumption i state =  | 
|
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1168  | 
  let val Thm{sign_ref,der,maxidx,hyps,prop,...} = state;
 | 
| 0 | 1169  | 
val (tpairs, Bs, Bi, C) = dest_state(state,i)  | 
| 250 | 1170  | 
      fun newth (env as Envir.Envir{maxidx, ...}, tpairs) =
 | 
| 1220 | 1171  | 
fix_shyps [state] (env_codT env)  | 
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1172  | 
          (Thm{sign_ref = sign_ref, 
 | 
| 2386 | 1173  | 
der = infer_derivs (Assumption (i, Some env), [der]),  | 
1174  | 
maxidx = maxidx,  | 
|
1175  | 
shyps = [],  | 
|
1176  | 
hyps = hyps,  | 
|
1177  | 
prop =  | 
|
1178  | 
if Envir.is_empty env then (*avoid wasted normalizations*)  | 
|
1179  | 
Logic.rule_of (tpairs, Bs, C)  | 
|
1180  | 
else (*normalize the new rule fully*)  | 
|
1181  | 
Envir.norm_term env (Logic.rule_of (tpairs, Bs, C))});  | 
|
| 0 | 1182  | 
fun addprfs [] = Sequence.null  | 
1183  | 
| addprfs ((t,u)::apairs) = Sequence.seqof (fn()=> Sequence.pull  | 
|
1184  | 
(Sequence.mapp newth  | 
|
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1185  | 
(Unify.unifiers(Sign.deref sign_ref,Envir.empty maxidx, (t,u)::tpairs))  | 
| 250 | 1186  | 
(addprfs apairs)))  | 
| 0 | 1187  | 
in addprfs (Logic.assum_pairs Bi) end;  | 
1188  | 
||
| 250 | 1189  | 
(*Solve subgoal Bi of proof state B1...Bn/C by assumption.  | 
| 0 | 1190  | 
Checks if Bi's conclusion is alpha-convertible to one of its assumptions*)  | 
1191  | 
fun eq_assumption i state =  | 
|
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1192  | 
  let val Thm{sign_ref,der,maxidx,hyps,prop,...} = state;
 | 
| 0 | 1193  | 
val (tpairs, Bs, Bi, C) = dest_state(state,i)  | 
1194  | 
in if exists (op aconv) (Logic.assum_pairs Bi)  | 
|
| 1220 | 1195  | 
then fix_shyps [state] []  | 
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1196  | 
             (Thm{sign_ref = sign_ref, 
 | 
| 2386 | 1197  | 
der = infer_derivs (Assumption (i,None), [der]),  | 
1198  | 
maxidx = maxidx,  | 
|
1199  | 
shyps = [],  | 
|
1200  | 
hyps = hyps,  | 
|
1201  | 
prop = Logic.rule_of(tpairs, Bs, C)})  | 
|
| 0 | 1202  | 
      else  raise THM("eq_assumption", 0, [state])
 | 
1203  | 
end;  | 
|
1204  | 
||
1205  | 
||
| 2671 | 1206  | 
(*For rotate_tac: fast rotation of assumptions of subgoal i*)  | 
1207  | 
fun rotate_rule k i state =  | 
|
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1208  | 
  let val Thm{sign_ref,der,maxidx,hyps,prop,shyps} = state;
 | 
| 2671 | 1209  | 
val (tpairs, Bs, Bi, C) = dest_state(state,i)  | 
1210  | 
val params = Logic.strip_params Bi  | 
|
1211  | 
and asms = Logic.strip_assums_hyp Bi  | 
|
1212  | 
and concl = Logic.strip_assums_concl Bi  | 
|
1213  | 
val n = length asms  | 
|
1214  | 
fun rot m = if 0=m orelse m=n then Bi  | 
|
1215  | 
else if 0<m andalso m<n  | 
|
1216  | 
then list_all  | 
|
1217  | 
(params,  | 
|
1218  | 
Logic.list_implies(List.drop(asms, m) @  | 
|
1219  | 
List.take(asms, m),  | 
|
1220  | 
concl))  | 
|
1221  | 
		   else raise THM("rotate_rule", m, [state])
 | 
|
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1222  | 
  in  Thm{sign_ref = sign_ref, 
 | 
| 2671 | 1223  | 
der = infer_derivs (Rotate_rule (k,i), [der]),  | 
1224  | 
maxidx = maxidx,  | 
|
1225  | 
shyps = shyps,  | 
|
1226  | 
hyps = hyps,  | 
|
1227  | 
prop = Logic.rule_of(tpairs, Bs@[rot (if k<0 then n+k else k)], C)}  | 
|
1228  | 
end;  | 
|
1229  | 
||
1230  | 
||
| 0 | 1231  | 
(** User renaming of parameters in a subgoal **)  | 
1232  | 
||
1233  | 
(*Calls error rather than raising an exception because it is intended  | 
|
1234  | 
for top-level use -- exception handling would not make sense here.  | 
|
1235  | 
The names in cs, if distinct, are used for the innermost parameters;  | 
|
1236  | 
preceding parameters may be renamed to make all params distinct.*)  | 
|
1237  | 
fun rename_params_rule (cs, i) state =  | 
|
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1238  | 
  let val Thm{sign_ref,der,maxidx,hyps,...} = state
 | 
| 0 | 1239  | 
val (tpairs, Bs, Bi, C) = dest_state(state,i)  | 
1240  | 
val iparams = map #1 (Logic.strip_params Bi)  | 
|
1241  | 
val short = length iparams - length cs  | 
|
| 250 | 1242  | 
val newnames =  | 
1243  | 
if short<0 then error"More names than abstractions!"  | 
|
1244  | 
else variantlist(take (short,iparams), cs) @ cs  | 
|
| 
3037
 
99ed078e6ae7
rename_params_rule used to check if the new name clashed with a free name in
 
nipkow 
parents: 
3012 
diff
changeset
 | 
1245  | 
val freenames = map (#1 o dest_Free) (term_frees Bi)  | 
| 0 | 1246  | 
val newBi = Logic.list_rename_params (newnames, Bi)  | 
| 250 | 1247  | 
in  | 
| 0 | 1248  | 
case findrep cs of  | 
| 
3565
 
c64410e701fb
Now rename_params_rule merely issues warnings--and does nothing--if the
 
paulson 
parents: 
3558 
diff
changeset
 | 
1249  | 
     c::_ => (warning ("Can't rename.  Bound variables not distinct: " ^ c); 
 | 
| 
 
c64410e701fb
Now rename_params_rule merely issues warnings--and does nothing--if the
 
paulson 
parents: 
3558 
diff
changeset
 | 
1250  | 
state)  | 
| 
1576
 
af8f43f742a0
Added some optimized versions of functions dealing with sets
 
berghofe 
parents: 
1569 
diff
changeset
 | 
1251  | 
| [] => (case cs inter_string freenames of  | 
| 
3565
 
c64410e701fb
Now rename_params_rule merely issues warnings--and does nothing--if the
 
paulson 
parents: 
3558 
diff
changeset
 | 
1252  | 
       a::_ => (warning ("Can't rename.  Bound/Free variable clash: " ^ a); 
 | 
| 
 
c64410e701fb
Now rename_params_rule merely issues warnings--and does nothing--if the
 
paulson 
parents: 
3558 
diff
changeset
 | 
1253  | 
state)  | 
| 1220 | 1254  | 
| [] => fix_shyps [state] []  | 
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1255  | 
                (Thm{sign_ref = sign_ref,
 | 
| 2386 | 1256  | 
der = infer_derivs (Rename_params_rule(cs,i), [der]),  | 
1257  | 
maxidx = maxidx,  | 
|
1258  | 
shyps = [],  | 
|
1259  | 
hyps = hyps,  | 
|
1260  | 
prop = Logic.rule_of(tpairs, Bs@[newBi], C)}))  | 
|
| 0 | 1261  | 
end;  | 
1262  | 
||
1263  | 
(*** Preservation of bound variable names ***)  | 
|
1264  | 
||
| 250 | 1265  | 
(*Scan a pair of terms; while they are similar,  | 
| 0 | 1266  | 
accumulate corresponding bound vars in "al"*)  | 
| 1238 | 1267  | 
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
 | 
1268  | 
match_bvs(s, t, if x="" orelse y="" then al  | 
| 1238 | 1269  | 
else (x,y)::al)  | 
| 0 | 1270  | 
| match_bvs(f$s, g$t, al) = match_bvs(f,g,match_bvs(s,t,al))  | 
1271  | 
| match_bvs(_,_,al) = al;  | 
|
1272  | 
||
1273  | 
(* strip abstractions created by parameters *)  | 
|
1274  | 
fun match_bvars((s,t),al) = match_bvs(strip_abs_body s, strip_abs_body t, al);  | 
|
1275  | 
||
1276  | 
||
| 250 | 1277  | 
(* strip_apply f A(,B) strips off all assumptions/parameters from A  | 
| 0 | 1278  | 
introduced by lifting over B, and applies f to remaining part of A*)  | 
1279  | 
fun strip_apply f =  | 
|
1280  | 
  let fun strip(Const("==>",_)$ A1 $ B1,
 | 
|
| 250 | 1281  | 
                Const("==>",_)$ _  $ B2) = implies $ A1 $ strip(B1,B2)
 | 
1282  | 
        | strip((c as Const("all",_)) $ Abs(a,T,t1),
 | 
|
1283  | 
                      Const("all",_)  $ Abs(_,_,t2)) = c$Abs(a,T,strip(t1,t2))
 | 
|
1284  | 
| strip(A,_) = f A  | 
|
| 0 | 1285  | 
in strip end;  | 
1286  | 
||
1287  | 
(*Use the alist to rename all bound variables and some unknowns in a term  | 
|
1288  | 
dpairs = current disagreement pairs; tpairs = permanent ones (flexflex);  | 
|
1289  | 
Preserves unknowns in tpairs and on lhs of dpairs. *)  | 
|
1290  | 
fun rename_bvs([],_,_,_) = I  | 
|
1291  | 
| rename_bvs(al,dpairs,tpairs,B) =  | 
|
| 250 | 1292  | 
let val vars = foldr add_term_vars  | 
1293  | 
(map fst dpairs @ map fst tpairs @ map snd tpairs, [])  | 
|
1294  | 
(*unknowns appearing elsewhere be preserved!*)  | 
|
1295  | 
val vids = map (#1 o #1 o dest_Var) vars;  | 
|
1296  | 
fun rename(t as Var((x,i),T)) =  | 
|
1297  | 
(case assoc(al,x) of  | 
|
| 
1576
 
af8f43f742a0
Added some optimized versions of functions dealing with sets
 
berghofe 
parents: 
1569 
diff
changeset
 | 
1298  | 
Some(y) => if x mem_string vids orelse y mem_string vids then t  | 
| 250 | 1299  | 
else Var((y,i),T)  | 
1300  | 
| None=> t)  | 
|
| 0 | 1301  | 
| rename(Abs(x,T,t)) =  | 
| 
1576
 
af8f43f742a0
Added some optimized versions of functions dealing with sets
 
berghofe 
parents: 
1569 
diff
changeset
 | 
1302  | 
Abs(case assoc_string(al,x) of Some(y) => y | None => x,  | 
| 250 | 1303  | 
T, rename t)  | 
| 0 | 1304  | 
| rename(f$t) = rename f $ rename t  | 
1305  | 
| rename(t) = t;  | 
|
| 250 | 1306  | 
fun strip_ren Ai = strip_apply rename (Ai,B)  | 
| 0 | 1307  | 
in strip_ren end;  | 
1308  | 
||
1309  | 
(*Function to rename bounds/unknowns in the argument, lifted over B*)  | 
|
1310  | 
fun rename_bvars(dpairs, tpairs, B) =  | 
|
| 250 | 1311  | 
rename_bvs(foldr match_bvars (dpairs,[]), dpairs, tpairs, B);  | 
| 0 | 1312  | 
|
1313  | 
||
1314  | 
(*** RESOLUTION ***)  | 
|
1315  | 
||
| 
721
 
479832ff2d29
Pure/thm/norm_term_skip: new, for skipping normalization of the context
 
lcp 
parents: 
678 
diff
changeset
 | 
1316  | 
(** Lifting optimizations **)  | 
| 
 
479832ff2d29
Pure/thm/norm_term_skip: new, for skipping normalization of the context
 
lcp 
parents: 
678 
diff
changeset
 | 
1317  | 
|
| 0 | 1318  | 
(*strip off pairs of assumptions/parameters in parallel -- they are  | 
1319  | 
identical because of lifting*)  | 
|
| 250 | 1320  | 
fun strip_assums2 (Const("==>", _) $ _ $ B1,
 | 
1321  | 
                   Const("==>", _) $ _ $ B2) = strip_assums2 (B1,B2)
 | 
|
| 0 | 1322  | 
  | strip_assums2 (Const("all",_)$Abs(a,T,t1),
 | 
| 250 | 1323  | 
                   Const("all",_)$Abs(_,_,t2)) =
 | 
| 0 | 1324  | 
let val (B1,B2) = strip_assums2 (t1,t2)  | 
1325  | 
in (Abs(a,T,B1), Abs(a,T,B2)) end  | 
|
1326  | 
| strip_assums2 BB = BB;  | 
|
1327  | 
||
1328  | 
||
| 
721
 
479832ff2d29
Pure/thm/norm_term_skip: new, for skipping normalization of the context
 
lcp 
parents: 
678 
diff
changeset
 | 
1329  | 
(*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
 | 
1330  | 
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
 | 
1331  | 
  | 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
 | 
1332  | 
        let val Envir.Envir{iTs, ...} = env
 | 
| 1238 | 1333  | 
val T' = typ_subst_TVars iTs T  | 
1334  | 
(*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
 | 
1335  | 
this could be a NEW parameter*)  | 
| 
 
479832ff2d29
Pure/thm/norm_term_skip: new, for skipping normalization of the context
 
lcp 
parents: 
678 
diff
changeset
 | 
1336  | 
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
 | 
1337  | 
  | norm_term_skip env n (Const("==>", _) $ A $ B) =
 | 
| 1238 | 1338  | 
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
 | 
1339  | 
| 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
 | 
1340  | 
|
| 
 
479832ff2d29
Pure/thm/norm_term_skip: new, for skipping normalization of the context
 
lcp 
parents: 
678 
diff
changeset
 | 
1341  | 
|
| 0 | 1342  | 
(*Composition of object rule r=(A1...Am/B) with proof state s=(B1...Bn/C)  | 
| 250 | 1343  | 
Unifies B with Bi, replacing subgoal i (1 <= i <= n)  | 
| 0 | 1344  | 
If match then forbid instantiations in proof state  | 
1345  | 
If lifted then shorten the dpair using strip_assums2.  | 
|
1346  | 
If eres_flg then simultaneously proves A1 by assumption.  | 
|
| 250 | 1347  | 
nsubgoal is the number of new subgoals (written m above).  | 
| 0 | 1348  | 
Curried so that resolution calls dest_state only once.  | 
1349  | 
*)  | 
|
| 1529 | 1350  | 
local open Sequence; exception COMPOSE  | 
| 0 | 1351  | 
in  | 
| 250 | 1352  | 
fun bicompose_aux match (state, (stpairs, Bs, Bi, C), lifted)  | 
| 0 | 1353  | 
(eres_flg, orule, nsubgoal) =  | 
| 1529 | 1354  | 
 let val Thm{der=sder, maxidx=smax, shyps=sshyps, hyps=shyps, ...} = state
 | 
1355  | 
     and Thm{der=rder, maxidx=rmax, shyps=rshyps, hyps=rhyps, 
 | 
|
| 2386 | 1356  | 
prop=rprop,...} = orule  | 
| 1529 | 1357  | 
(*How many hyps to skip over during normalization*)  | 
| 1238 | 1358  | 
and nlift = Logic.count_prems(strip_all_body Bi,  | 
1359  | 
if eres_flg then ~1 else 0)  | 
|
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1360  | 
val sign_ref = merge_thm_sgs(state,orule);  | 
| 
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1361  | 
val sign = Sign.deref sign_ref;  | 
| 0 | 1362  | 
(** Add new theorem with prop = '[| Bs; As |] ==> C' to thq **)  | 
| 250 | 1363  | 
     fun addth As ((env as Envir.Envir {maxidx, ...}, tpairs), thq) =
 | 
1364  | 
let val normt = Envir.norm_term env;  | 
|
1365  | 
(*perform minimal copying here by examining env*)  | 
|
1366  | 
val normp =  | 
|
1367  | 
if Envir.is_empty env then (tpairs, Bs @ As, C)  | 
|
1368  | 
else  | 
|
1369  | 
let val ntps = map (pairself normt) tpairs  | 
|
| 2147 | 1370  | 
in if Envir.above (smax, env) then  | 
| 1238 | 1371  | 
(*no assignments in state; normalize the rule only*)  | 
1372  | 
if lifted  | 
|
1373  | 
then (ntps, Bs @ map (norm_term_skip env nlift) As, C)  | 
|
1374  | 
else (ntps, Bs @ map normt As, C)  | 
|
| 1529 | 1375  | 
else if match then raise COMPOSE  | 
| 250 | 1376  | 
else (*normalize the new rule fully*)  | 
1377  | 
(ntps, map normt (Bs @ As), normt C)  | 
|
1378  | 
end  | 
|
| 1258 | 1379  | 
val th = (*tuned fix_shyps*)  | 
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1380  | 
             Thm{sign_ref = sign_ref,
 | 
| 2386 | 1381  | 
der = infer_derivs (Bicompose(match, eres_flg,  | 
1382  | 
1 + length Bs, nsubgoal, env),  | 
|
1383  | 
[rder,sder]),  | 
|
1384  | 
maxidx = maxidx,  | 
|
1385  | 
shyps = add_env_sorts (env, union_sort(rshyps,sshyps)),  | 
|
1386  | 
hyps = union_term(rhyps,shyps),  | 
|
1387  | 
prop = Logic.rule_of normp}  | 
|
| 1529 | 1388  | 
in cons(th, thq) end handle COMPOSE => thq  | 
| 0 | 1389  | 
val (rtpairs,rhorn) = Logic.strip_flexpairs(rprop);  | 
1390  | 
val (rAs,B) = Logic.strip_prems(nsubgoal, [], rhorn)  | 
|
1391  | 
       handle TERM _ => raise THM("bicompose: rule", 0, [orule,state]);
 | 
|
1392  | 
(*Modify assumptions, deleting n-th if n>0 for e-resolution*)  | 
|
1393  | 
fun newAs(As0, n, dpairs, tpairs) =  | 
|
1394  | 
let val As1 = if !Logic.auto_rename orelse not lifted then As0  | 
|
| 250 | 1395  | 
else map (rename_bvars(dpairs,tpairs,B)) As0  | 
| 0 | 1396  | 
in (map (Logic.flatten_params n) As1)  | 
| 250 | 1397  | 
handle TERM _ =>  | 
1398  | 
          raise THM("bicompose: 1st premise", 0, [orule])
 | 
|
| 0 | 1399  | 
end;  | 
| 2147 | 1400  | 
val env = Envir.empty(Int.max(rmax,smax));  | 
| 0 | 1401  | 
val BBi = if lifted then strip_assums2(B,Bi) else (B,Bi);  | 
1402  | 
val dpairs = BBi :: (rtpairs@stpairs);  | 
|
1403  | 
(*elim-resolution: try each assumption in turn. Initially n=1*)  | 
|
1404  | 
fun tryasms (_, _, []) = null  | 
|
1405  | 
| tryasms (As, n, (t,u)::apairs) =  | 
|
| 250 | 1406  | 
(case pull(Unify.unifiers(sign, env, (t,u)::dpairs)) of  | 
1407  | 
None => tryasms (As, n+1, apairs)  | 
|
1408  | 
| cell as Some((_,tpairs),_) =>  | 
|
1409  | 
its_right (addth (newAs(As, n, [BBi,(u,t)], tpairs)))  | 
|
1410  | 
(seqof (fn()=> cell),  | 
|
1411  | 
seqof (fn()=> pull (tryasms (As, n+1, apairs)))));  | 
|
| 0 | 1412  | 
     fun eres [] = raise THM("bicompose: no premises", 0, [orule,state])
 | 
1413  | 
| eres (A1::As) = tryasms (As, 1, Logic.assum_pairs A1);  | 
|
1414  | 
(*ordinary resolution*)  | 
|
1415  | 
fun res(None) = null  | 
|
| 250 | 1416  | 
| res(cell as Some((_,tpairs),_)) =  | 
1417  | 
its_right (addth(newAs(rev rAs, 0, [BBi], tpairs)))  | 
|
1418  | 
(seqof (fn()=> cell), null)  | 
|
| 0 | 1419  | 
in if eres_flg then eres(rev rAs)  | 
1420  | 
else res(pull(Unify.unifiers(sign, env, dpairs)))  | 
|
1421  | 
end;  | 
|
1422  | 
end; (*open Sequence*)  | 
|
1423  | 
||
1424  | 
||
1425  | 
fun bicompose match arg i state =  | 
|
1426  | 
bicompose_aux match (state, dest_state(state,i), false) arg;  | 
|
1427  | 
||
1428  | 
(*Quick test whether rule is resolvable with the subgoal with hyps Hs  | 
|
1429  | 
and conclusion B. If eres_flg then checks 1st premise of rule also*)  | 
|
1430  | 
fun could_bires (Hs, B, eres_flg, rule) =  | 
|
1431  | 
let fun could_reshyp (A1::_) = exists (apl(A1,could_unify)) Hs  | 
|
| 250 | 1432  | 
| could_reshyp [] = false; (*no premise -- illegal*)  | 
1433  | 
in could_unify(concl_of rule, B) andalso  | 
|
1434  | 
(not eres_flg orelse could_reshyp (prems_of rule))  | 
|
| 0 | 1435  | 
end;  | 
1436  | 
||
1437  | 
(*Bi-resolution of a state with a list of (flag,rule) pairs.  | 
|
1438  | 
Puts the rule above: rule/state. Renames vars in the rules. *)  | 
|
| 250 | 1439  | 
fun biresolution match brules i state =  | 
| 0 | 1440  | 
let val lift = lift_rule(state, i);  | 
| 250 | 1441  | 
val (stpairs, Bs, Bi, C) = dest_state(state,i)  | 
1442  | 
val B = Logic.strip_assums_concl Bi;  | 
|
1443  | 
val Hs = Logic.strip_assums_hyp Bi;  | 
|
1444  | 
val comp = bicompose_aux match (state, (stpairs, Bs, Bi, C), true);  | 
|
1445  | 
fun res [] = Sequence.null  | 
|
1446  | 
| res ((eres_flg, rule)::brules) =  | 
|
1447  | 
if could_bires (Hs, B, eres_flg, rule)  | 
|
| 1160 | 1448  | 
then Sequence.seqof (*delay processing remainder till needed*)  | 
| 250 | 1449  | 
(fn()=> Some(comp (eres_flg, lift rule, nprems_of rule),  | 
1450  | 
res brules))  | 
|
1451  | 
else res brules  | 
|
| 0 | 1452  | 
in Sequence.flats (res brules) end;  | 
1453  | 
||
1454  | 
||
1455  | 
||
| 2509 | 1456  | 
(*** Meta Simplification ***)  | 
| 0 | 1457  | 
|
| 2509 | 1458  | 
(** diagnostics **)  | 
| 0 | 1459  | 
|
1460  | 
exception SIMPLIFIER of string * thm;  | 
|
1461  | 
||
| 2509 | 1462  | 
fun prtm a sign t = (writeln a; writeln (Sign.string_of_term sign t));  | 
| 3529 | 1463  | 
fun prtm_warning a sign t = (warning a; warning (Sign.string_of_term sign t));  | 
| 
1580
 
e3fd931e6095
Added some functions which allow redirection of Isabelle's output
 
berghofe 
parents: 
1576 
diff
changeset
 | 
1464  | 
|
| 209 | 1465  | 
val trace_simp = ref false;  | 
1466  | 
||
| 3558 | 1467  | 
fun trace a = if ! trace_simp then writeln a else ();  | 
| 2509 | 1468  | 
fun trace_warning a = if ! trace_simp then warning a else ();  | 
1469  | 
fun trace_term a sign t = if ! trace_simp then prtm a sign t else ();  | 
|
1470  | 
fun trace_term_warning a sign t = if ! trace_simp then prtm_warning a sign t else ();  | 
|
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1471  | 
|
| 
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1472  | 
fun trace_thm a (Thm {sign_ref, prop, ...}) =
 | 
| 
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1473  | 
trace_term a (Sign.deref sign_ref) prop;  | 
| 
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1474  | 
|
| 
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1475  | 
fun trace_thm_warning a (Thm {sign_ref, prop, ...}) =
 | 
| 
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1476  | 
trace_term_warning a (Sign.deref sign_ref) prop;  | 
| 209 | 1477  | 
|
1478  | 
||
| 
1580
 
e3fd931e6095
Added some functions which allow redirection of Isabelle's output
 
berghofe 
parents: 
1576 
diff
changeset
 | 
1479  | 
|
| 2509 | 1480  | 
(** meta simp sets **)  | 
1481  | 
||
1482  | 
(* basic components *)  | 
|
| 
1580
 
e3fd931e6095
Added some functions which allow redirection of Isabelle's output
 
berghofe 
parents: 
1576 
diff
changeset
 | 
1483  | 
|
| 2509 | 1484  | 
type rrule = {thm: thm, lhs: term, perm: bool};
 | 
1485  | 
type cong = {thm: thm, lhs: term};
 | 
|
| 
3577
 
9715b6e3ec5f
added prems argument to simplification procedures;
 
wenzelm 
parents: 
3565 
diff
changeset
 | 
1486  | 
type simproc =  | 
| 
 
9715b6e3ec5f
added prems argument to simplification procedures;
 
wenzelm 
parents: 
3565 
diff
changeset
 | 
1487  | 
 {name: string, proc: Sign.sg -> thm list -> term -> thm option, lhs: cterm, id: stamp};
 | 
| 
288
 
b00ce6a1fe27
Implemented "ordered rewriting": rules which merely permute variables, such
 
nipkow 
parents: 
274 
diff
changeset
 | 
1488  | 
|
| 3550 | 1489  | 
fun eq_rrule ({thm = Thm {prop = p1, ...}, ...}: rrule,
 | 
| 2509 | 1490  | 
  {thm = Thm {prop = p2, ...}, ...}: rrule) = p1 aconv p2;
 | 
1491  | 
||
| 3550 | 1492  | 
fun eq_cong ({thm = Thm {prop = p1, ...}, ...}: cong,
 | 
1493  | 
  {thm = Thm {prop = p2, ...}, ...}: cong) = p1 aconv p2;
 | 
|
1494  | 
||
1495  | 
fun eq_prem (Thm {prop = p1, ...}, Thm {prop = p2, ...}) = p1 aconv p2;
 | 
|
1496  | 
||
1497  | 
fun eq_simproc ({id = s1, ...}:simproc, {id = s2, ...}:simproc) = (s1 = s2);
 | 
|
1498  | 
||
1499  | 
fun mk_simproc (name, proc, lhs, id) =  | 
|
1500  | 
  {name = name, proc = proc, lhs = lhs, id = id};
 | 
|
| 2509 | 1501  | 
|
1502  | 
||
1503  | 
(* datatype mss *)  | 
|
| 
288
 
b00ce6a1fe27
Implemented "ordered rewriting": rules which merely permute variables, such
 
nipkow 
parents: 
274 
diff
changeset
 | 
1504  | 
|
| 2509 | 1505  | 
(*  | 
1506  | 
A "mss" contains data needed during conversion:  | 
|
1507  | 
rules: discrimination net of rewrite rules;  | 
|
1508  | 
congs: association list of congruence rules;  | 
|
1509  | 
procs: discrimination net of simplification procedures  | 
|
1510  | 
(functions that prove rewrite rules on the fly);  | 
|
1511  | 
bounds: names of bound variables already used  | 
|
1512  | 
(for generating new names when rewriting under lambda abstractions);  | 
|
1513  | 
prems: current premises;  | 
|
1514  | 
mk_rews: turns simplification thms into rewrite rules;  | 
|
1515  | 
termless: relation for ordered rewriting;  | 
|
| 1028 | 1516  | 
*)  | 
| 0 | 1517  | 
|
| 2509 | 1518  | 
datatype meta_simpset =  | 
1519  | 
  Mss of {
 | 
|
1520  | 
rules: rrule Net.net,  | 
|
1521  | 
congs: (string * cong) list,  | 
|
1522  | 
procs: simproc Net.net,  | 
|
1523  | 
bounds: string list,  | 
|
1524  | 
prems: thm list,  | 
|
1525  | 
mk_rews: thm -> thm list,  | 
|
1526  | 
termless: term * term -> bool};  | 
|
1527  | 
||
1528  | 
fun mk_mss (rules, congs, procs, bounds, prems, mk_rews, termless) =  | 
|
1529  | 
  Mss {rules = rules, congs = congs, procs = procs, bounds = bounds,
 | 
|
1530  | 
prems = prems, mk_rews = mk_rews, termless = termless};  | 
|
1531  | 
||
1532  | 
val empty_mss =  | 
|
1533  | 
mk_mss (Net.empty, [], Net.empty, [], [], K [], Logic.termless);  | 
|
1534  | 
||
1535  | 
||
1536  | 
||
1537  | 
(** simpset operations **)  | 
|
1538  | 
||
| 3550 | 1539  | 
(* dest_mss *)  | 
1540  | 
||
1541  | 
fun dest_mss (Mss {rules, congs, procs, ...}) =
 | 
|
1542  | 
  {simps = map (fn (_, {thm, ...}) => thm) (Net.dest rules),
 | 
|
1543  | 
   congs = map (fn (_, {thm, ...}) => thm) congs,
 | 
|
1544  | 
procs =  | 
|
1545  | 
     map (fn (_, {name, lhs, id, ...}) => ((name, lhs), id)) (Net.dest procs)
 | 
|
1546  | 
|> partition_eq eq_snd  | 
|
1547  | 
|> map (fn ps => (#1 (#1 (hd ps)), map (#2 o #1) ps))};  | 
|
1548  | 
||
1549  | 
||
1550  | 
(* merge_mss *) (*NOTE: ignores mk_rews and termless of 2nd mss*)  | 
|
1551  | 
||
1552  | 
fun merge_mss  | 
|
1553  | 
 (Mss {rules = rules1, congs = congs1, procs = procs1, bounds = bounds1,
 | 
|
1554  | 
prems = prems1, mk_rews, termless},  | 
|
1555  | 
  Mss {rules = rules2, congs = congs2, procs = procs2, bounds = bounds2,
 | 
|
1556  | 
prems = prems2, ...}) =  | 
|
1557  | 
mk_mss  | 
|
1558  | 
(Net.merge (rules1, rules2, eq_rrule),  | 
|
1559  | 
generic_merge (eq_cong o pairself snd) I I congs1 congs2,  | 
|
1560  | 
Net.merge (procs1, procs2, eq_simproc),  | 
|
1561  | 
merge_lists bounds1 bounds2,  | 
|
1562  | 
generic_merge eq_prem I I prems1 prems2,  | 
|
1563  | 
mk_rews, termless);  | 
|
1564  | 
||
1565  | 
||
| 2509 | 1566  | 
(* mk_rrule *)  | 
1567  | 
||
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1568  | 
fun mk_rrule (thm as Thm {sign_ref, prop, ...}) =
 | 
| 1238 | 1569  | 
let  | 
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1570  | 
val sign = Sign.deref sign_ref;  | 
| 2509 | 1571  | 
val prems = Logic.strip_imp_prems prop;  | 
1572  | 
val concl = Logic.strip_imp_concl prop;  | 
|
| 
3893
 
5a1f22e7b359
The simplifier has been improved a little: equations s=t which used to be
 
nipkow 
parents: 
3812 
diff
changeset
 | 
1573  | 
val (lhs, rhs) = Logic.dest_equals concl handle TERM _ =>  | 
| 2509 | 1574  | 
      raise SIMPLIFIER ("Rewrite rule not a meta-equality", thm);
 | 
| 
3893
 
5a1f22e7b359
The simplifier has been improved a little: equations s=t which used to be
 
nipkow 
parents: 
3812 
diff
changeset
 | 
1575  | 
in case Logic.loops sign prems lhs rhs of  | 
| 
 
5a1f22e7b359
The simplifier has been improved a little: equations s=t which used to be
 
nipkow 
parents: 
3812 
diff
changeset
 | 
1576  | 
     (None,perm) => Some {thm = thm, lhs = lhs, perm = perm}
 | 
| 
 
5a1f22e7b359
The simplifier has been improved a little: equations s=t which used to be
 
nipkow 
parents: 
3812 
diff
changeset
 | 
1577  | 
| (Some msg,_) =>  | 
| 
 
5a1f22e7b359
The simplifier has been improved a little: equations s=t which used to be
 
nipkow 
parents: 
3812 
diff
changeset
 | 
1578  | 
        (prtm_warning("ignoring rewrite rule ("^msg^")") sign prop; None)
 | 
| 0 | 1579  | 
end;  | 
1580  | 
||
| 2509 | 1581  | 
|
1582  | 
(* add_simps *)  | 
|
| 87 | 1583  | 
|
| 2509 | 1584  | 
fun add_simp  | 
1585  | 
  (mss as Mss {rules, congs, procs, bounds, prems, mk_rews, termless},
 | 
|
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1586  | 
    thm as Thm {sign_ref, prop, ...}) =
 | 
| 2509 | 1587  | 
(case mk_rrule thm of  | 
| 87 | 1588  | 
None => mss  | 
| 2509 | 1589  | 
  | Some (rrule as {lhs, ...}) =>
 | 
| 209 | 1590  | 
(trace_thm "Adding rewrite rule:" thm;  | 
| 2509 | 1591  | 
mk_mss (Net.insert_term ((lhs, rrule), rules, eq_rrule) handle Net.INSERT =>  | 
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1592  | 
(prtm_warning "ignoring duplicate rewrite rule" (Sign.deref sign_ref) prop; rules),  | 
| 2509 | 1593  | 
congs, procs, bounds, prems, mk_rews, termless)));  | 
| 0 | 1594  | 
|
1595  | 
val add_simps = foldl add_simp;  | 
|
| 2509 | 1596  | 
|
1597  | 
fun mss_of thms = add_simps (empty_mss, thms);  | 
|
1598  | 
||
1599  | 
||
1600  | 
(* del_simps *)  | 
|
1601  | 
||
1602  | 
fun del_simp  | 
|
1603  | 
  (mss as Mss {rules, congs, procs, bounds, prems, mk_rews, termless},
 | 
|
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1604  | 
    thm as Thm {sign_ref, prop, ...}) =
 | 
| 2509 | 1605  | 
(case mk_rrule thm of  | 
1606  | 
None => mss  | 
|
1607  | 
  | Some (rrule as {lhs, ...}) =>
 | 
|
1608  | 
mk_mss (Net.delete_term ((lhs, rrule), rules, eq_rrule) handle Net.DELETE =>  | 
|
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1609  | 
(prtm_warning "rewrite rule not in simpset" (Sign.deref sign_ref) prop; rules),  | 
| 2509 | 1610  | 
congs, procs, bounds, prems, mk_rews, termless));  | 
1611  | 
||
| 87 | 1612  | 
val del_simps = foldl del_simp;  | 
| 0 | 1613  | 
|
| 2509 | 1614  | 
|
| 2626 | 1615  | 
(* add_congs *)  | 
| 0 | 1616  | 
|
| 2509 | 1617  | 
fun add_cong (Mss {rules, congs, procs, bounds, prems, mk_rews, termless}, thm) =
 | 
1618  | 
let  | 
|
1619  | 
val (lhs, _) = Logic.dest_equals (concl_of thm) handle TERM _ =>  | 
|
1620  | 
      raise SIMPLIFIER ("Congruence not a meta-equality", thm);
 | 
|
1621  | 
(* val lhs = Pattern.eta_contract lhs; *)  | 
|
1622  | 
val (a, _) = dest_Const (head_of lhs) handle TERM _ =>  | 
|
1623  | 
      raise SIMPLIFIER ("Congruence must start with a constant", thm);
 | 
|
1624  | 
in  | 
|
1625  | 
    mk_mss (rules, (a, {lhs = lhs, thm = thm}) :: congs, procs, bounds,
 | 
|
1626  | 
prems, mk_rews, termless)  | 
|
| 0 | 1627  | 
end;  | 
1628  | 
||
1629  | 
val (op add_congs) = foldl add_cong;  | 
|
1630  | 
||
| 2509 | 1631  | 
|
| 2626 | 1632  | 
(* del_congs *)  | 
1633  | 
||
1634  | 
fun del_cong (Mss {rules, congs, procs, bounds, prems, mk_rews, termless}, thm) =
 | 
|
1635  | 
let  | 
|
1636  | 
val (lhs, _) = Logic.dest_equals (concl_of thm) handle TERM _ =>  | 
|
1637  | 
      raise SIMPLIFIER ("Congruence not a meta-equality", thm);
 | 
|
1638  | 
(* val lhs = Pattern.eta_contract lhs; *)  | 
|
1639  | 
val (a, _) = dest_Const (head_of lhs) handle TERM _ =>  | 
|
1640  | 
      raise SIMPLIFIER ("Congruence must start with a constant", thm);
 | 
|
1641  | 
in  | 
|
1642  | 
mk_mss (rules, filter (fn (x,_)=> x<>a) congs, procs, bounds,  | 
|
1643  | 
prems, mk_rews, termless)  | 
|
1644  | 
end;  | 
|
1645  | 
||
1646  | 
val (op del_congs) = foldl del_cong;  | 
|
1647  | 
||
1648  | 
||
| 2509 | 1649  | 
(* add_simprocs *)  | 
1650  | 
||
| 3550 | 1651  | 
fun add_proc (mss as Mss {rules, congs, procs, bounds, prems, mk_rews, termless},
 | 
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1652  | 
    (name, lhs as Cterm {sign_ref, t, ...}, proc, id)) =
 | 
| 
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1653  | 
  (trace_term ("Adding simplification procedure " ^ name ^ " for:")
 | 
| 
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1654  | 
(Sign.deref sign_ref) t;  | 
| 2509 | 1655  | 
mk_mss (rules, congs,  | 
| 3550 | 1656  | 
Net.insert_term ((t, mk_simproc (name, proc, lhs, id)), procs, eq_simproc)  | 
1657  | 
handle Net.INSERT => (trace_warning "ignored duplicate"; procs),  | 
|
| 2509 | 1658  | 
bounds, prems, mk_rews, termless));  | 
| 0 | 1659  | 
|
| 3550 | 1660  | 
fun add_simproc (mss, (name, lhss, proc, id)) =  | 
1661  | 
foldl add_proc (mss, map (fn lhs => (name, lhs, proc, id)) lhss);  | 
|
1662  | 
||
| 2509 | 1663  | 
val add_simprocs = foldl add_simproc;  | 
1664  | 
||
1665  | 
||
1666  | 
(* del_simprocs *)  | 
|
| 0 | 1667  | 
|
| 3550 | 1668  | 
fun del_proc (mss as Mss {rules, congs, procs, bounds, prems, mk_rews, termless},
 | 
1669  | 
    (name, lhs as Cterm {t, ...}, proc, id)) =
 | 
|
| 2509 | 1670  | 
mk_mss (rules, congs,  | 
| 3550 | 1671  | 
Net.delete_term ((t, mk_simproc (name, proc, lhs, id)), procs, eq_simproc)  | 
1672  | 
handle Net.DELETE => (trace_warning "simplification procedure not in simpset"; procs),  | 
|
1673  | 
bounds, prems, mk_rews, termless);  | 
|
1674  | 
||
1675  | 
fun del_simproc (mss, (name, lhss, proc, id)) =  | 
|
1676  | 
foldl del_proc (mss, map (fn lhs => (name, lhs, proc, id)) lhss);  | 
|
| 2509 | 1677  | 
|
1678  | 
val del_simprocs = foldl del_simproc;  | 
|
| 0 | 1679  | 
|
1680  | 
||
| 2509 | 1681  | 
(* prems *)  | 
1682  | 
||
1683  | 
fun add_prems (Mss {rules, congs, procs, bounds, prems, mk_rews, termless}, thms) =
 | 
|
1684  | 
mk_mss (rules, congs, procs, bounds, thms @ prems, mk_rews, termless);  | 
|
1685  | 
||
1686  | 
fun prems_of_mss (Mss {prems, ...}) = prems;
 | 
|
1687  | 
||
1688  | 
||
1689  | 
(* mk_rews *)  | 
|
1690  | 
||
1691  | 
fun set_mk_rews  | 
|
1692  | 
  (Mss {rules, congs, procs, bounds, prems, mk_rews = _, termless}, mk_rews) =
 | 
|
1693  | 
mk_mss (rules, congs, procs, bounds, prems, mk_rews, termless);  | 
|
1694  | 
||
1695  | 
fun mk_rews_of_mss (Mss {mk_rews, ...}) = mk_rews;
 | 
|
1696  | 
||
1697  | 
||
1698  | 
(* termless *)  | 
|
1699  | 
||
1700  | 
fun set_termless  | 
|
1701  | 
  (Mss {rules, congs, procs, bounds, prems, mk_rews, termless = _}, termless) =
 | 
|
1702  | 
mk_mss (rules, congs, procs, bounds, prems, mk_rews, termless);  | 
|
1703  | 
||
1704  | 
||
1705  | 
||
1706  | 
(** rewriting **)  | 
|
1707  | 
||
1708  | 
(*  | 
|
1709  | 
Uses conversions, omitting proofs for efficiency. See:  | 
|
1710  | 
L C Paulson, A higher-order implementation of rewriting,  | 
|
1711  | 
Science of Computer Programming 3 (1983), pages 119-149.  | 
|
1712  | 
*)  | 
|
| 0 | 1713  | 
|
1714  | 
type prover = meta_simpset -> thm -> thm option;  | 
|
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1715  | 
type termrec = (Sign.sg_ref * term list) * term;  | 
| 0 | 1716  | 
type conv = meta_simpset -> termrec -> termrec;  | 
1717  | 
||
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1718  | 
fun check_conv (thm as Thm{shyps,hyps,prop,sign_ref,der,maxidx,...}, prop0, ders) =
 | 
| 432 | 1719  | 
let fun err() = (trace_thm "Proved wrong thm (Check subgoaler?)" thm;  | 
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1720  | 
trace_term "Should have proved" (Sign.deref sign_ref) prop0;  | 
| 432 | 1721  | 
None)  | 
| 0 | 1722  | 
val (lhs0,_) = Logic.dest_equals(Logic.strip_imp_concl prop0)  | 
1723  | 
in case prop of  | 
|
1724  | 
       Const("==",_) $ lhs $ rhs =>
 | 
|
1725  | 
if (lhs = lhs0) orelse  | 
|
| 
427
 
4ce2ce940faf
ordered rewriting applies to conditional rules as well now
 
nipkow 
parents: 
421 
diff
changeset
 | 
1726  | 
(lhs aconv Envir.norm_term (Envir.empty 0) lhs0)  | 
| 1529 | 1727  | 
then (trace_thm "SUCCEEDED" thm;  | 
| 2386 | 1728  | 
Some(shyps, hyps, maxidx, rhs, der::ders))  | 
| 0 | 1729  | 
else err()  | 
1730  | 
| _ => err()  | 
|
1731  | 
end;  | 
|
1732  | 
||
| 
659
 
95ed2ccb4438
Prepared the code for preservation of bound var names during rewriting. Still
 
nipkow 
parents: 
623 
diff
changeset
 | 
1733  | 
fun ren_inst(insts,prop,pat,obj) =  | 
| 
 
95ed2ccb4438
Prepared the code for preservation of bound var names during rewriting. Still
 
nipkow 
parents: 
623 
diff
changeset
 | 
1734  | 
let val ren = match_bvs(pat,obj,[])  | 
| 
 
95ed2ccb4438
Prepared the code for preservation of bound var names during rewriting. Still
 
nipkow 
parents: 
623 
diff
changeset
 | 
1735  | 
fun renAbs(Abs(x,T,b)) =  | 
| 
1576
 
af8f43f742a0
Added some optimized versions of functions dealing with sets
 
berghofe 
parents: 
1569 
diff
changeset
 | 
1736  | 
Abs(case assoc_string(ren,x) of None => x | Some(y) => y, T, renAbs(b))  | 
| 
659
 
95ed2ccb4438
Prepared the code for preservation of bound var names during rewriting. Still
 
nipkow 
parents: 
623 
diff
changeset
 | 
1737  | 
| renAbs(f$t) = renAbs(f) $ renAbs(t)  | 
| 
 
95ed2ccb4438
Prepared the code for preservation of bound var names during rewriting. Still
 
nipkow 
parents: 
623 
diff
changeset
 | 
1738  | 
| renAbs(t) = t  | 
| 
 
95ed2ccb4438
Prepared the code for preservation of bound var names during rewriting. Still
 
nipkow 
parents: 
623 
diff
changeset
 | 
1739  | 
in subst_vars insts (if null(ren) then prop else renAbs(prop)) end;  | 
| 
678
 
6151b7f3b606
Modified pattern.ML to perform proper matching of Higher-Order Patterns.
 
nipkow 
parents: 
659 
diff
changeset
 | 
1740  | 
|
| 1258 | 1741  | 
fun add_insts_sorts ((iTs, is), Ss) =  | 
1742  | 
add_typs_sorts (map snd iTs, add_terms_sorts (map snd is, Ss));  | 
|
1743  | 
||
| 
659
 
95ed2ccb4438
Prepared the code for preservation of bound var names during rewriting. Still
 
nipkow 
parents: 
623 
diff
changeset
 | 
1744  | 
|
| 2509 | 1745  | 
(* mk_procrule *)  | 
1746  | 
||
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1747  | 
fun mk_procrule (thm as Thm {sign_ref, prop, ...}) =
 | 
| 2509 | 1748  | 
let  | 
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1749  | 
val sign = Sign.deref sign_ref;  | 
| 2509 | 1750  | 
val prems = Logic.strip_imp_prems prop;  | 
1751  | 
val concl = Logic.strip_imp_concl prop;  | 
|
1752  | 
val (lhs, _) = Logic.dest_equals concl handle TERM _ =>  | 
|
1753  | 
      raise SIMPLIFIER ("Rewrite rule not a meta-equality", thm);
 | 
|
1754  | 
val econcl = Pattern.eta_contract concl;  | 
|
1755  | 
val (elhs, erhs) = Logic.dest_equals econcl;  | 
|
1756  | 
in  | 
|
1757  | 
if not ((term_vars erhs) subset  | 
|
| 2671 | 1758  | 
(union_term (term_vars elhs, List.concat(map term_vars prems))))  | 
1759  | 
then (prtm_warning "extra Var(s) on rhs" sign prop; [])  | 
|
| 2509 | 1760  | 
    else [{thm = thm, lhs = lhs, perm = false}]
 | 
1761  | 
end;  | 
|
1762  | 
||
1763  | 
||
1764  | 
(* conversion to apply the meta simpset to a term *)  | 
|
1765  | 
||
1766  | 
(*  | 
|
1767  | 
we try in order:  | 
|
1768  | 
(1) beta reduction  | 
|
1769  | 
(2) unconditional rewrite rules  | 
|
1770  | 
(3) conditional rewrite rules  | 
|
| 3550 | 1771  | 
(4) simplification procedures  | 
| 2509 | 1772  | 
*)  | 
1773  | 
||
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1774  | 
fun rewritec (prover,sign_reft) (mss as Mss{rules, procs, mk_rews, termless, prems, ...}) 
 | 
| 2147 | 1775  | 
(shypst,hypst,maxt,t,ders) =  | 
| 3550 | 1776  | 
let  | 
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1777  | 
val signt = Sign.deref sign_reft;  | 
| 
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1778  | 
val tsigt = Sign.tsig_of signt;  | 
| 
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1779  | 
      fun rew {thm as Thm{sign_ref,der,maxidx,shyps,hyps,prop,...}, lhs, perm} =
 | 
| 
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1780  | 
let  | 
| 
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1781  | 
val _ =  | 
| 
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1782  | 
if Sign.subsig (Sign.deref sign_ref, signt) then ()  | 
| 
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1783  | 
else (trace_thm_warning "rewrite rule from different theory" thm;  | 
| 
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1784  | 
raise Pattern.MATCH);  | 
| 2147 | 1785  | 
val rprop = if maxt = ~1 then prop  | 
1786  | 
else Logic.incr_indexes([],maxt+1) prop;  | 
|
1787  | 
val rlhs = if maxt = ~1 then lhs  | 
|
| 
1065
 
8425cb5acb77
Fixed old bug in the simplifier. Term to be simplified now carries around its
 
nipkow 
parents: 
1028 
diff
changeset
 | 
1788  | 
else fst(Logic.dest_equals(Logic.strip_imp_concl rprop))  | 
| 3550 | 1789  | 
val insts = Pattern.match tsigt (rlhs,t);  | 
| 
1065
 
8425cb5acb77
Fixed old bug in the simplifier. Term to be simplified now carries around its
 
nipkow 
parents: 
1028 
diff
changeset
 | 
1790  | 
val prop' = ren_inst(insts,rprop,rlhs,t);  | 
| 
2177
 
8b365a3a6ed1
Changed some mem, ins and union calls to be monomorphic
 
paulson 
parents: 
2163 
diff
changeset
 | 
1791  | 
val hyps' = union_term(hyps,hypst);  | 
| 
 
8b365a3a6ed1
Changed some mem, ins and union calls to be monomorphic
 
paulson 
parents: 
2163 
diff
changeset
 | 
1792  | 
val shyps' = add_insts_sorts (insts, union_sort(shyps,shypst));  | 
| 
1065
 
8425cb5acb77
Fixed old bug in the simplifier. Term to be simplified now carries around its
 
nipkow 
parents: 
1028 
diff
changeset
 | 
1793  | 
val maxidx' = maxidx_of_term prop'  | 
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1794  | 
            val ct' = Cterm{sign_ref = sign_reft,       (*used for deriv only*)
 | 
| 2386 | 1795  | 
t = prop',  | 
1796  | 
T = propT,  | 
|
1797  | 
maxidx = maxidx'}  | 
|
| 3550 | 1798  | 
val der' = infer_derivs (RewriteC ct', [der]);  | 
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1799  | 
            val thm' = Thm{sign_ref = sign_reft, 
 | 
| 2386 | 1800  | 
der = der',  | 
1801  | 
shyps = shyps',  | 
|
1802  | 
hyps = hyps',  | 
|
| 1529 | 1803  | 
prop = prop',  | 
| 2386 | 1804  | 
maxidx = maxidx'}  | 
| 
427
 
4ce2ce940faf
ordered rewriting applies to conditional rules as well now
 
nipkow 
parents: 
421 
diff
changeset
 | 
1805  | 
val (lhs',rhs') = Logic.dest_equals(Logic.strip_imp_concl prop')  | 
| 
 
4ce2ce940faf
ordered rewriting applies to conditional rules as well now
 
nipkow 
parents: 
421 
diff
changeset
 | 
1806  | 
in if perm andalso not(termless(rhs',lhs')) then None else  | 
| 
 
4ce2ce940faf
ordered rewriting applies to conditional rules as well now
 
nipkow 
parents: 
421 
diff
changeset
 | 
1807  | 
if Logic.count_prems(prop',0) = 0  | 
| 1529 | 1808  | 
then (trace_thm "Rewriting:" thm';  | 
| 2386 | 1809  | 
Some(shyps', hyps', maxidx', rhs', der'::ders))  | 
| 0 | 1810  | 
else (trace_thm "Trying to rewrite:" thm';  | 
1811  | 
case prover mss thm' of  | 
|
1812  | 
None => (trace_thm "FAILED" thm'; None)  | 
|
| 1529 | 1813  | 
| Some(thm2) => check_conv(thm2,prop',ders))  | 
| 0 | 1814  | 
end  | 
1815  | 
||
| 
225
 
76f60e6400e8
optimized net for matching of abstractions to speed up simplifier
 
nipkow 
parents: 
222 
diff
changeset
 | 
1816  | 
fun rews [] = None  | 
| 2509 | 1817  | 
| rews (rrule :: rrules) =  | 
| 
225
 
76f60e6400e8
optimized net for matching of abstractions to speed up simplifier
 
nipkow 
parents: 
222 
diff
changeset
 | 
1818  | 
let val opt = rew rrule handle Pattern.MATCH => None  | 
| 
 
76f60e6400e8
optimized net for matching of abstractions to speed up simplifier
 
nipkow 
parents: 
222 
diff
changeset
 | 
1819  | 
in case opt of None => rews rrules | some => some end;  | 
| 3550 | 1820  | 
|
| 1659 | 1821  | 
fun sort_rrules rrs = let  | 
| 2386 | 1822  | 
        fun is_simple {thm as Thm{prop,...}, lhs, perm} = case prop of 
 | 
1823  | 
                                        Const("==",_) $ _ $ _ => true
 | 
|
1824  | 
| _ => false  | 
|
1825  | 
fun sort [] (re1,re2) = re1 @ re2  | 
|
1826  | 
| sort (rr::rrs) (re1,re2) = if is_simple rr  | 
|
1827  | 
then sort rrs (rr::re1,re2)  | 
|
1828  | 
else sort rrs (re1,rr::re2)  | 
|
| 1659 | 1829  | 
in sort rrs ([],[])  | 
1830  | 
end  | 
|
| 2509 | 1831  | 
|
| 3550 | 1832  | 
fun proc_rews _ ([]:simproc list) = None  | 
1833  | 
        | proc_rews eta_t ({name, proc, lhs = Cterm {t = plhs, ...}, ...} :: ps) =
 | 
|
1834  | 
if Pattern.matches tsigt (plhs, t) then  | 
|
| 3558 | 1835  | 
             (trace_term ("Trying procedure " ^ name ^ " on:") signt eta_t;
 | 
| 
3577
 
9715b6e3ec5f
added prems argument to simplification procedures;
 
wenzelm 
parents: 
3565 
diff
changeset
 | 
1836  | 
case proc signt prems eta_t of  | 
| 3558 | 1837  | 
None => (trace "FAILED"; proc_rews eta_t ps)  | 
| 3550 | 1838  | 
| Some raw_thm =>  | 
| 3558 | 1839  | 
                 (trace_thm ("Procedure " ^ name ^ " proved rewrite rule:") raw_thm;
 | 
| 3550 | 1840  | 
(case rews (mk_procrule raw_thm) of  | 
| 3558 | 1841  | 
None => (trace "IGNORED"; proc_rews eta_t ps)  | 
| 3550 | 1842  | 
| some => some)))  | 
1843  | 
else proc_rews eta_t ps;  | 
|
| 2509 | 1844  | 
in  | 
| 2792 | 1845  | 
(case t of  | 
| 3550 | 1846  | 
Abs (_, _, body) $ u =>  | 
| 2509 | 1847  | 
Some (shypst, hypst, maxt, subst_bound (u, body), ders)  | 
1848  | 
| _ =>  | 
|
| 2792 | 1849  | 
(case rews (sort_rrules (Net.match_term rules t)) of  | 
| 3012 | 1850  | 
None => proc_rews (Pattern.eta_contract t) (Net.match_term procs t)  | 
| 2509 | 1851  | 
| some => some))  | 
| 0 | 1852  | 
end;  | 
1853  | 
||
| 2509 | 1854  | 
|
1855  | 
(* conversion to apply a congruence rule to a term *)  | 
|
1856  | 
||
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1857  | 
fun congc (prover,sign_reft) {thm=cong,lhs=lhs} (shypst,hypst,maxt,t,ders) =
 | 
| 
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1858  | 
let val signt = Sign.deref sign_reft;  | 
| 
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1859  | 
val tsig = Sign.tsig_of signt;  | 
| 
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1860  | 
      val Thm{sign_ref,der,shyps,hyps,maxidx,prop,...} = cong
 | 
| 
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1861  | 
val _ = if Sign.subsig(Sign.deref sign_ref,signt) then ()  | 
| 
208
 
342f88d2e8ab
optimized simplifier - signature of rewritten term stays constant
 
nipkow 
parents: 
200 
diff
changeset
 | 
1862  | 
                 else error("Congruence rule from different theory")
 | 
| 2147 | 1863  | 
val rprop = if maxt = ~1 then prop  | 
1864  | 
else Logic.incr_indexes([],maxt+1) prop;  | 
|
1865  | 
val rlhs = if maxt = ~1 then lhs  | 
|
| 
1065
 
8425cb5acb77
Fixed old bug in the simplifier. Term to be simplified now carries around its
 
nipkow 
parents: 
1028 
diff
changeset
 | 
1866  | 
else fst(Logic.dest_equals(Logic.strip_imp_concl rprop))  | 
| 
1569
 
a89f246dee0e
Non-matching congruence rule in rewriter is simply ignored.
 
nipkow 
parents: 
1566 
diff
changeset
 | 
1867  | 
val insts = Pattern.match tsig (rlhs,t)  | 
| 
 
a89f246dee0e
Non-matching congruence rule in rewriter is simply ignored.
 
nipkow 
parents: 
1566 
diff
changeset
 | 
1868  | 
(* Pattern.match can raise Pattern.MATCH;  | 
| 
 
a89f246dee0e
Non-matching congruence rule in rewriter is simply ignored.
 
nipkow 
parents: 
1566 
diff
changeset
 | 
1869  | 
is handled when congc is called *)  | 
| 
1065
 
8425cb5acb77
Fixed old bug in the simplifier. Term to be simplified now carries around its
 
nipkow 
parents: 
1028 
diff
changeset
 | 
1870  | 
val prop' = ren_inst(insts,rprop,rlhs,t);  | 
| 
2177
 
8b365a3a6ed1
Changed some mem, ins and union calls to be monomorphic
 
paulson 
parents: 
2163 
diff
changeset
 | 
1871  | 
val shyps' = add_insts_sorts (insts, union_sort(shyps,shypst))  | 
| 1529 | 1872  | 
val maxidx' = maxidx_of_term prop'  | 
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1873  | 
      val ct' = Cterm{sign_ref = sign_reft,     (*used for deriv only*)
 | 
| 2386 | 1874  | 
t = prop',  | 
1875  | 
T = propT,  | 
|
1876  | 
maxidx = maxidx'}  | 
|
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1877  | 
      val thm' = Thm{sign_ref = sign_reft, 
 | 
| 3550 | 1878  | 
der = infer_derivs (CongC ct', [der]),  | 
| 2386 | 1879  | 
shyps = shyps',  | 
1880  | 
hyps = union_term(hyps,hypst),  | 
|
| 1529 | 1881  | 
prop = prop',  | 
| 2386 | 1882  | 
maxidx = maxidx'};  | 
| 0 | 1883  | 
val unit = trace_thm "Applying congruence rule" thm';  | 
| 
112
 
009ae5c85ae9
Changed the simplifier: if the subgoaler proves an unexpected thm, chances
 
nipkow 
parents: 
87 
diff
changeset
 | 
1884  | 
      fun err() = error("Failed congruence proof!")
 | 
| 0 | 1885  | 
|
1886  | 
in case prover thm' of  | 
|
| 
112
 
009ae5c85ae9
Changed the simplifier: if the subgoaler proves an unexpected thm, chances
 
nipkow 
parents: 
87 
diff
changeset
 | 
1887  | 
None => err()  | 
| 1529 | 1888  | 
| Some(thm2) => (case check_conv(thm2,prop',ders) of  | 
| 
405
 
c97514f63633
Internale optimization of the simplifier: in case a subterm stays unchanged,
 
nipkow 
parents: 
399 
diff
changeset
 | 
1889  | 
None => err() | some => some)  | 
| 0 | 1890  | 
end;  | 
1891  | 
||
1892  | 
||
| 
405
 
c97514f63633
Internale optimization of the simplifier: in case a subterm stays unchanged,
 
nipkow 
parents: 
399 
diff
changeset
 | 
1893  | 
|
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1894  | 
fun bottomc ((simprem,useprem),prover,sign_ref) =  | 
| 1529 | 1895  | 
let fun botc fail mss trec =  | 
| 2386 | 1896  | 
(case subc mss trec of  | 
1897  | 
some as Some(trec1) =>  | 
|
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1898  | 
(case rewritec (prover,sign_ref) mss trec1 of  | 
| 2386 | 1899  | 
Some(trec2) => botc false mss trec2  | 
1900  | 
| None => some)  | 
|
1901  | 
| None =>  | 
|
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1902  | 
(case rewritec (prover,sign_ref) mss trec of  | 
| 2386 | 1903  | 
Some(trec2) => botc false mss trec2  | 
1904  | 
| None => if fail then None else Some(trec)))  | 
|
| 0 | 1905  | 
|
| 1529 | 1906  | 
and try_botc mss trec = (case botc true mss trec of  | 
| 2386 | 1907  | 
Some(trec1) => trec1  | 
1908  | 
| None => trec)  | 
|
| 
405
 
c97514f63633
Internale optimization of the simplifier: in case a subterm stays unchanged,
 
nipkow 
parents: 
399 
diff
changeset
 | 
1909  | 
|
| 2509 | 1910  | 
     and subc (mss as Mss{rules,congs,procs,bounds,prems,mk_rews,termless})
 | 
| 2386 | 1911  | 
(trec as (shyps,hyps,maxidx,t0,ders)) =  | 
| 1529 | 1912  | 
(case t0 of  | 
| 2386 | 1913  | 
Abs(a,T,t) =>  | 
1914  | 
let val b = variant bounds a  | 
|
1915  | 
                 val v = Free("." ^ b,T)
 | 
|
| 2509 | 1916  | 
val mss' = mk_mss (rules, congs, procs, b :: bounds, prems, mk_rews, termless)  | 
| 2386 | 1917  | 
in case botc true mss'  | 
1918  | 
(shyps,hyps,maxidx,subst_bound (v,t),ders) of  | 
|
1919  | 
Some(shyps',hyps',maxidx',t',ders') =>  | 
|
1920  | 
Some(shyps', hyps', maxidx',  | 
|
1921  | 
Abs(a, T, abstract_over(v,t')),  | 
|
1922  | 
ders')  | 
|
1923  | 
| None => None  | 
|
1924  | 
end  | 
|
1925  | 
| t$u => (case t of  | 
|
1926  | 
             Const("==>",_)$s  => Some(impc(shyps,hyps,maxidx,s,u,mss,ders))
 | 
|
1927  | 
| Abs(_,_,body) =>  | 
|
1928  | 
let val trec = (shyps,hyps,maxidx,subst_bound (u,body),ders)  | 
|
1929  | 
in case subc mss trec of  | 
|
1930  | 
None => Some(trec)  | 
|
1931  | 
| trec => trec  | 
|
1932  | 
end  | 
|
1933  | 
| _ =>  | 
|
1934  | 
let fun appc() =  | 
|
1935  | 
(case botc true mss (shyps,hyps,maxidx,t,ders) of  | 
|
1936  | 
Some(shyps1,hyps1,maxidx1,t1,ders1) =>  | 
|
1937  | 
(case botc true mss (shyps1,hyps1,maxidx,u,ders1) of  | 
|
1938  | 
Some(shyps2,hyps2,maxidx2,u1,ders2) =>  | 
|
1939  | 
Some(shyps2, hyps2, Int.max(maxidx1,maxidx2),  | 
|
1940  | 
t1$u1, ders2)  | 
|
1941  | 
| None =>  | 
|
1942  | 
Some(shyps1, hyps1, Int.max(maxidx1,maxidx), t1$u,  | 
|
1943  | 
ders1))  | 
|
1944  | 
| None =>  | 
|
1945  | 
(case botc true mss (shyps,hyps,maxidx,u,ders) of  | 
|
1946  | 
Some(shyps1,hyps1,maxidx1,u1,ders1) =>  | 
|
1947  | 
Some(shyps1, hyps1, Int.max(maxidx,maxidx1),  | 
|
1948  | 
t$u1, ders1)  | 
|
1949  | 
| None => None))  | 
|
1950  | 
val (h,ts) = strip_comb t  | 
|
1951  | 
in case h of  | 
|
1952  | 
Const(a,_) =>  | 
|
1953  | 
(case assoc_string(congs,a) of  | 
|
1954  | 
None => appc()  | 
|
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1955  | 
| Some(cong) => (congc (prover mss,sign_ref) cong trec  | 
| 
1569
 
a89f246dee0e
Non-matching congruence rule in rewriter is simply ignored.
 
nipkow 
parents: 
1566 
diff
changeset
 | 
1956  | 
handle Pattern.MATCH => appc() ) )  | 
| 2386 | 1957  | 
| _ => appc()  | 
1958  | 
end)  | 
|
1959  | 
| _ => None)  | 
|
| 0 | 1960  | 
|
| 1529 | 1961  | 
     and impc(shyps, hyps, maxidx, s, u, mss as Mss{mk_rews,...}, ders) =
 | 
1962  | 
let val (shyps1,hyps1,_,s1,ders1) =  | 
|
| 2386 | 1963  | 
if simprem then try_botc mss (shyps,hyps,maxidx,s,ders)  | 
1964  | 
else (shyps,hyps,0,s,ders);  | 
|
1965  | 
val maxidx1 = maxidx_of_term s1  | 
|
1966  | 
val mss1 =  | 
|
| 
2535
 
907a3379f165
Added warning msg when the simplifier cannot use a premise as a rewrite rule
 
nipkow 
parents: 
2509 
diff
changeset
 | 
1967  | 
if not useprem then mss else  | 
| 
2620
 
54f21bf9522a
Made troublesome simplifier warning dependent on trace_simp.
 
nipkow 
parents: 
2535 
diff
changeset
 | 
1968  | 
if maxidx1 <> ~1 then (trace_term_warning  | 
| 
2535
 
907a3379f165
Added warning msg when the simplifier cannot use a premise as a rewrite rule
 
nipkow 
parents: 
2509 
diff
changeset
 | 
1969  | 
"Cannot add premise as rewrite rule because it contains (type) unknowns:"  | 
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1970  | 
(Sign.deref sign_ref) s1; mss)  | 
| 
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1971  | 
             else let val thm = assume (Cterm{sign_ref=sign_ref, t=s1, 
 | 
| 2386 | 1972  | 
T=propT, maxidx=maxidx1})  | 
1973  | 
in add_simps(add_prems(mss,[thm]), mk_rews thm) end  | 
|
1974  | 
val (shyps2,hyps2,maxidx2,u1,ders2) =  | 
|
1975  | 
try_botc mss1 (shyps1,hyps1,maxidx,u,ders1)  | 
|
1976  | 
val hyps3 = if gen_mem (op aconv) (s1, hyps1)  | 
|
1977  | 
then hyps2 else hyps2\s1  | 
|
| 2147 | 1978  | 
in (shyps2, hyps3, Int.max(maxidx1,maxidx2),  | 
| 2386 | 1979  | 
Logic.mk_implies(s1,u1), ders2)  | 
| 1529 | 1980  | 
end  | 
| 0 | 1981  | 
|
| 1529 | 1982  | 
in try_botc end;  | 
| 0 | 1983  | 
|
1984  | 
||
1985  | 
(*** Meta-rewriting: rewrites t to u and returns the theorem t==u ***)  | 
|
| 2509 | 1986  | 
|
1987  | 
(*  | 
|
1988  | 
Parameters:  | 
|
1989  | 
mode = (simplify A, use A in simplifying B) when simplifying A ==> B  | 
|
1990  | 
mss: contains equality theorems of the form [|p1,...|] ==> t==u  | 
|
1991  | 
prover: how to solve premises in conditional rewrites and congruences  | 
|
| 0 | 1992  | 
*)  | 
| 2509 | 1993  | 
|
1994  | 
(* FIXME: check that #bounds(mss) does not "occur" in ct alread *)  | 
|
1995  | 
||
| 
214
 
ed6a3e2b1a33
added new parameter to the simplification tactics which indicates if
 
nipkow 
parents: 
209 
diff
changeset
 | 
1996  | 
fun rewrite_cterm mode mss prover ct =  | 
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1997  | 
  let val Cterm {sign_ref, t, T, maxidx} = ct;
 | 
| 2147 | 1998  | 
val (shyps,hyps,maxu,u,ders) =  | 
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
1999  | 
bottomc (mode,prover, sign_ref) mss  | 
| 2386 | 2000  | 
(add_term_sorts(t,[]), [], maxidx, t, []);  | 
| 0 | 2001  | 
val prop = Logic.mk_equals(t,u)  | 
| 1258 | 2002  | 
in  | 
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
2003  | 
      Thm{sign_ref = sign_ref, 
 | 
| 2386 | 2004  | 
der = infer_derivs (Rewrite_cterm ct, ders),  | 
2005  | 
maxidx = Int.max (maxidx,maxu),  | 
|
2006  | 
shyps = shyps,  | 
|
2007  | 
hyps = hyps,  | 
|
| 1529 | 2008  | 
prop = prop}  | 
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
2009  | 
end;  | 
| 0 | 2010  | 
|
| 1539 | 2011  | 
|
| 2509 | 2012  | 
|
2013  | 
(*** Oracles ***)  | 
|
2014  | 
||
| 3812 | 2015  | 
fun invoke_oracle thy raw_name =  | 
2016  | 
let  | 
|
2017  | 
    val {sign = sg, oracles, ...} = rep_theory thy;
 | 
|
2018  | 
val name = Sign.intern sg Theory.oracleK raw_name;  | 
|
2019  | 
val oracle =  | 
|
2020  | 
(case Symtab.lookup (oracles, name) of  | 
|
2021  | 
        None => raise THM ("Unknown oracle: " ^ name, 0, [])
 | 
|
2022  | 
| Some (f, _) => f);  | 
|
2023  | 
in  | 
|
2024  | 
fn (sign, exn) =>  | 
|
2025  | 
let  | 
|
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
2026  | 
val sign_ref' = Sign.merge_refs (Sign.self_ref sg, Sign.self_ref sign);  | 
| 
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
2027  | 
val sign' = Sign.deref sign_ref';  | 
| 3812 | 2028  | 
val (prop, T, maxidx) = Sign.certify_term sign' (oracle (sign', exn));  | 
2029  | 
in  | 
|
2030  | 
if T <> propT then  | 
|
2031  | 
          raise THM ("Oracle's result must have type prop: " ^ name, 0, [])
 | 
|
2032  | 
else fix_shyps [] []  | 
|
| 
3967
 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 
wenzelm 
parents: 
3895 
diff
changeset
 | 
2033  | 
          (Thm {sign_ref = sign_ref', 
 | 
| 3812 | 2034  | 
der = Join (Oracle (thy, name, sign, exn), []),  | 
2035  | 
maxidx = maxidx,  | 
|
2036  | 
shyps = [],  | 
|
2037  | 
hyps = [],  | 
|
2038  | 
prop = prop})  | 
|
2039  | 
end  | 
|
2040  | 
end;  | 
|
2041  | 
||
| 1539 | 2042  | 
|
| 0 | 2043  | 
end;  | 
| 1503 | 2044  | 
|
2045  | 
open Thm;  |