| author | paulson | 
| Fri, 23 Dec 2005 17:34:46 +0100 | |
| changeset 18508 | c5861e128a95 | 
| parent 18498 | 466351242c6f | 
| child 18535 | 84b0597808bb | 
| permissions | -rw-r--r-- | 
| 
252
 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 
wenzelm 
parents: 
229 
diff
changeset
 | 
1  | 
(* Title: Pure/drule.ML  | 
| 0 | 2  | 
ID: $Id$  | 
| 
252
 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 
wenzelm 
parents: 
229 
diff
changeset
 | 
3  | 
Author: Lawrence C Paulson, Cambridge University Computer Laboratory  | 
| 0 | 4  | 
Copyright 1993 University of Cambridge  | 
5  | 
||
| 3766 | 6  | 
Derived rules and other operations on theorems.  | 
| 0 | 7  | 
*)  | 
8  | 
||
| 
13606
 
2f121149acfe
Removed nRS again because extract_rews now takes care of preserving names.
 
berghofe 
parents: 
13569 
diff
changeset
 | 
9  | 
infix 0 RS RSN RL RLN MRS MRL OF COMP;  | 
| 0 | 10  | 
|
| 5903 | 11  | 
signature BASIC_DRULE =  | 
| 3766 | 12  | 
sig  | 
| 18179 | 13  | 
val mk_implies: cterm * cterm -> cterm  | 
14  | 
val list_implies: cterm list * cterm -> cterm  | 
|
15  | 
val dest_implies: cterm -> cterm * cterm  | 
|
16  | 
val dest_equals: cterm -> cterm * cterm  | 
|
17  | 
val strip_imp_prems: cterm -> cterm list  | 
|
18  | 
val strip_imp_concl: cterm -> cterm  | 
|
19  | 
val cprems_of: thm -> cterm list  | 
|
20  | 
val cterm_fun: (term -> term) -> (cterm -> cterm)  | 
|
21  | 
val ctyp_fun: (typ -> typ) -> (ctyp -> ctyp)  | 
|
| 18206 | 22  | 
val read_insts: theory -> (indexname -> typ option) * (indexname -> sort option) ->  | 
23  | 
(indexname -> typ option) * (indexname -> sort option) -> string list ->  | 
|
24  | 
(indexname * string) list -> (ctyp * ctyp) list * (cterm * cterm) list  | 
|
| 4285 | 25  | 
val types_sorts: thm -> (indexname-> typ option) * (indexname-> sort option)  | 
| 18179 | 26  | 
val strip_shyps_warning: thm -> thm  | 
27  | 
val forall_intr_list: cterm list -> thm -> thm  | 
|
28  | 
val forall_intr_frees: thm -> thm  | 
|
29  | 
val forall_intr_vars: thm -> thm  | 
|
30  | 
val forall_elim_list: cterm list -> thm -> thm  | 
|
31  | 
val forall_elim_var: int -> thm -> thm  | 
|
32  | 
val forall_elim_vars: int -> thm -> thm  | 
|
33  | 
val gen_all: thm -> thm  | 
|
34  | 
val lift_all: cterm -> thm -> thm  | 
|
35  | 
val freeze_thaw: thm -> thm * (thm -> thm)  | 
|
| 
15495
 
50fde6f04e4c
new treatment of demodulation in proof reconstruction
 
paulson 
parents: 
15442 
diff
changeset
 | 
36  | 
val freeze_thaw_robust: thm -> thm * (int -> thm -> thm)  | 
| 18179 | 37  | 
val implies_elim_list: thm -> thm list -> thm  | 
38  | 
val implies_intr_list: cterm list -> thm -> thm  | 
|
| 18206 | 39  | 
val instantiate: (ctyp * ctyp) list * (cterm * cterm) list -> thm -> thm  | 
| 18179 | 40  | 
val zero_var_indexes: thm -> thm  | 
41  | 
val implies_intr_hyps: thm -> thm  | 
|
42  | 
val standard: thm -> thm  | 
|
43  | 
val standard': thm -> thm  | 
|
44  | 
val rotate_prems: int -> thm -> thm  | 
|
45  | 
val rearrange_prems: int list -> thm -> thm  | 
|
46  | 
val assume_ax: theory -> string -> thm  | 
|
47  | 
val RSN: thm * (int * thm) -> thm  | 
|
48  | 
val RS: thm * thm -> thm  | 
|
49  | 
val RLN: thm list * (int * thm list) -> thm list  | 
|
50  | 
val RL: thm list * thm list -> thm list  | 
|
51  | 
val MRS: thm list * thm -> thm  | 
|
52  | 
val MRL: thm list list * thm list -> thm list  | 
|
53  | 
val OF: thm * thm list -> thm  | 
|
54  | 
val compose: thm * int * thm -> thm list  | 
|
55  | 
val COMP: thm * thm -> thm  | 
|
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
15949 
diff
changeset
 | 
56  | 
val read_instantiate_sg: theory -> (string*string)list -> thm -> thm  | 
| 18179 | 57  | 
val read_instantiate: (string*string)list -> thm -> thm  | 
58  | 
val cterm_instantiate: (cterm*cterm)list -> thm -> thm  | 
|
59  | 
val eq_thm_thy: thm * thm -> bool  | 
|
60  | 
val eq_thm_prop: thm * thm -> bool  | 
|
61  | 
val weak_eq_thm: thm * thm -> bool  | 
|
62  | 
val size_of_thm: thm -> int  | 
|
63  | 
val reflexive_thm: thm  | 
|
64  | 
val symmetric_thm: thm  | 
|
65  | 
val transitive_thm: thm  | 
|
66  | 
val symmetric_fun: thm -> thm  | 
|
67  | 
val extensional: thm -> thm  | 
|
68  | 
val imp_cong: thm  | 
|
69  | 
val swap_prems_eq: thm  | 
|
70  | 
val equal_abs_elim: cterm -> thm -> thm  | 
|
| 4285 | 71  | 
val equal_abs_elim_list: cterm list -> thm -> thm  | 
| 18179 | 72  | 
val asm_rl: thm  | 
73  | 
val cut_rl: thm  | 
|
74  | 
val revcut_rl: thm  | 
|
75  | 
val thin_rl: thm  | 
|
| 4285 | 76  | 
val triv_forall_equality: thm  | 
| 18179 | 77  | 
val swap_prems_rl: thm  | 
78  | 
val equal_intr_rule: thm  | 
|
79  | 
val equal_elim_rule1: thm  | 
|
80  | 
val inst: string -> string -> thm -> thm  | 
|
81  | 
val instantiate': ctyp option list -> cterm option list -> thm -> thm  | 
|
82  | 
val incr_indexes: thm -> thm -> thm  | 
|
83  | 
val incr_indexes_wrt: int list -> ctyp list -> cterm list -> thm list -> thm -> thm  | 
|
| 5903 | 84  | 
end;  | 
85  | 
||
86  | 
signature DRULE =  | 
|
87  | 
sig  | 
|
88  | 
include BASIC_DRULE  | 
|
| 15949 | 89  | 
val list_comb: cterm * cterm list -> cterm  | 
| 
12908
 
53bfe07a7916
New function strip_comb (cterm version of Term.strip_comb).
 
berghofe 
parents: 
12800 
diff
changeset
 | 
90  | 
val strip_comb: cterm -> cterm * cterm list  | 
| 15262 | 91  | 
val strip_type: ctyp -> ctyp list * ctyp  | 
| 15949 | 92  | 
val beta_conv: cterm -> cterm -> cterm  | 
| 15875 | 93  | 
val plain_prop_of: thm -> term  | 
| 15669 | 94  | 
val add_used: thm -> string list -> string list  | 
| 11975 | 95  | 
  val rule_attribute: ('a -> thm -> thm) -> 'a attribute
 | 
| 18225 | 96  | 
val map_tags: (tag list -> tag list) -> thm -> thm  | 
| 11975 | 97  | 
val tag_rule: tag -> thm -> thm  | 
98  | 
val untag_rule: string -> thm -> thm  | 
|
99  | 
val tag: tag -> 'a attribute  | 
|
100  | 
val untag: string -> 'a attribute  | 
|
101  | 
val get_kind: thm -> string  | 
|
102  | 
val kind: string -> 'a attribute  | 
|
103  | 
val theoremK: string  | 
|
104  | 
val lemmaK: string  | 
|
105  | 
val corollaryK: string  | 
|
106  | 
val internalK: string  | 
|
107  | 
val kind_internal: 'a attribute  | 
|
108  | 
val has_internal: tag list -> bool  | 
|
| 18468 | 109  | 
val is_internal: thm -> bool  | 
| 17713 | 110  | 
val flexflex_unique: thm -> thm  | 
| 11975 | 111  | 
val close_derivation: thm -> thm  | 
| 12005 | 112  | 
val local_standard: thm -> thm  | 
| 11975 | 113  | 
val compose_single: thm * int * thm -> thm  | 
| 12373 | 114  | 
val add_rule: thm -> thm list -> thm list  | 
115  | 
val del_rule: thm -> thm list -> thm list  | 
|
| 11975 | 116  | 
val add_rules: thm list -> thm list -> thm list  | 
117  | 
val del_rules: thm list -> thm list -> thm list  | 
|
118  | 
val merge_rules: thm list * thm list -> thm list  | 
|
| 18468 | 119  | 
val imp_cong_rule: thm -> thm -> thm  | 
| 
15001
 
fb2141a9f8c0
Moved conversion rules from MetaSimplifier to Drule. refl_implies removed
 
skalberg 
parents: 
14854 
diff
changeset
 | 
120  | 
val beta_eta_conversion: cterm -> thm  | 
| 15925 | 121  | 
val eta_long_conversion: cterm -> thm  | 
| 18468 | 122  | 
val forall_conv: int -> (cterm -> thm) -> cterm -> thm  | 
123  | 
val concl_conv: int -> (cterm -> thm) -> cterm -> thm  | 
|
124  | 
val prems_conv: int -> (int -> cterm -> thm) -> cterm -> thm  | 
|
125  | 
val conjunction_conv: int -> (int -> cterm -> thm) -> cterm -> thm  | 
|
| 18179 | 126  | 
val goals_conv: (int -> bool) -> (cterm -> thm) -> cterm -> thm  | 
127  | 
val fconv_rule: (cterm -> thm) -> thm -> thm  | 
|
| 11975 | 128  | 
val norm_hhf_eq: thm  | 
| 12800 | 129  | 
val is_norm_hhf: term -> bool  | 
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
15949 
diff
changeset
 | 
130  | 
val norm_hhf: theory -> term -> term  | 
| 18025 | 131  | 
val protect: cterm -> cterm  | 
132  | 
val protectI: thm  | 
|
133  | 
val protectD: thm  | 
|
| 18179 | 134  | 
val protect_cong: thm  | 
| 18025 | 135  | 
val implies_intr_protected: cterm list -> thm -> thm  | 
| 11975 | 136  | 
val freeze_all: thm -> thm  | 
137  | 
val tvars_of_terms: term list -> (indexname * sort) list  | 
|
138  | 
val vars_of_terms: term list -> (indexname * typ) list  | 
|
139  | 
val tvars_of: thm -> (indexname * sort) list  | 
|
140  | 
val vars_of: thm -> (indexname * typ) list  | 
|
| 18129 | 141  | 
val tfrees_of: thm -> (string * sort) list  | 
142  | 
val frees_of: thm -> (string * typ) list  | 
|
143  | 
val fold_terms: (term -> 'a -> 'a) -> thm -> 'a -> 'a  | 
|
| 14081 | 144  | 
val rename_bvars: (string * string) list -> thm -> thm  | 
145  | 
val rename_bvars': string option list -> thm -> thm  | 
|
| 11975 | 146  | 
val unvarifyT: thm -> thm  | 
147  | 
val unvarify: thm -> thm  | 
|
| 18129 | 148  | 
val tvars_intr_list: string list -> thm -> (string * (indexname * sort)) list * thm  | 
| 12297 | 149  | 
val remdups_rl: thm  | 
| 18225 | 150  | 
val multi_resolve: thm list -> thm -> thm Seq.seq  | 
151  | 
val multi_resolves: thm list -> thm list -> thm Seq.seq  | 
|
| 11975 | 152  | 
val conj_intr: thm -> thm -> thm  | 
153  | 
val conj_intr_list: thm list -> thm  | 
|
154  | 
val conj_elim: thm -> thm * thm  | 
|
155  | 
val conj_elim_list: thm -> thm list  | 
|
| 
18498
 
466351242c6f
conj_elim_precise: proper treatment of nested conjunctions;
 
wenzelm 
parents: 
18468 
diff
changeset
 | 
156  | 
val conj_elim_precise: int list -> thm -> thm list list  | 
| 12135 | 157  | 
val conj_intr_thm: thm  | 
| 18206 | 158  | 
val conj_curry: thm -> thm  | 
| 13325 | 159  | 
val abs_def: thm -> thm  | 
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
15949 
diff
changeset
 | 
160  | 
val read_instantiate_sg': theory -> (indexname * string) list -> thm -> thm  | 
| 15797 | 161  | 
val read_instantiate': (indexname * string) list -> thm -> thm  | 
| 3766 | 162  | 
end;  | 
| 0 | 163  | 
|
| 5903 | 164  | 
structure Drule: DRULE =  | 
| 0 | 165  | 
struct  | 
166  | 
||
| 3991 | 167  | 
|
| 16682 | 168  | 
(** some cterm->cterm operations: faster than calling cterm_of! **)  | 
| 
708
 
8422e50adce0
Pure/drule/cprems_of, cskip_flexpairs, cstrip_imp_prems: new cterm operations
 
lcp 
parents: 
668 
diff
changeset
 | 
169  | 
|
| 
2004
 
3411fe560611
New operations on cterms.  Now same names as in Logic
 
paulson 
parents: 
1906 
diff
changeset
 | 
170  | 
fun dest_implies ct =  | 
| 16682 | 171  | 
(case Thm.term_of ct of  | 
172  | 
    (Const ("==>", _) $ _ $ _) =>
 | 
|
173  | 
let val (ct1, ct2) = Thm.dest_comb ct  | 
|
174  | 
in (#2 (Thm.dest_comb ct1), ct2) end  | 
|
175  | 
  | _ => raise TERM ("dest_implies", [term_of ct]));
 | 
|
| 
1703
 
e22ad43bab5f
moved dest_cimplies to drule.ML; added adjust_maxidx
 
clasohm 
parents: 
1596 
diff
changeset
 | 
176  | 
|
| 10414 | 177  | 
fun dest_equals ct =  | 
| 16682 | 178  | 
(case Thm.term_of ct of  | 
179  | 
    (Const ("==", _) $ _ $ _) =>
 | 
|
180  | 
let val (ct1, ct2) = Thm.dest_comb ct  | 
|
181  | 
in (#2 (Thm.dest_comb ct1), ct2) end  | 
|
182  | 
    | _ => raise TERM ("dest_equals", [term_of ct]));
 | 
|
| 10414 | 183  | 
|
| 
1703
 
e22ad43bab5f
moved dest_cimplies to drule.ML; added adjust_maxidx
 
clasohm 
parents: 
1596 
diff
changeset
 | 
184  | 
|
| 
708
 
8422e50adce0
Pure/drule/cprems_of, cskip_flexpairs, cstrip_imp_prems: new cterm operations
 
lcp 
parents: 
668 
diff
changeset
 | 
185  | 
(* A1==>...An==>B goes to [A1,...,An], where B is not an implication *)  | 
| 
2004
 
3411fe560611
New operations on cterms.  Now same names as in Logic
 
paulson 
parents: 
1906 
diff
changeset
 | 
186  | 
fun strip_imp_prems ct =  | 
| 
 
3411fe560611
New operations on cterms.  Now same names as in Logic
 
paulson 
parents: 
1906 
diff
changeset
 | 
187  | 
let val (cA,cB) = dest_implies ct  | 
| 
 
3411fe560611
New operations on cterms.  Now same names as in Logic
 
paulson 
parents: 
1906 
diff
changeset
 | 
188  | 
in cA :: strip_imp_prems cB end  | 
| 
708
 
8422e50adce0
Pure/drule/cprems_of, cskip_flexpairs, cstrip_imp_prems: new cterm operations
 
lcp 
parents: 
668 
diff
changeset
 | 
189  | 
handle TERM _ => [];  | 
| 
 
8422e50adce0
Pure/drule/cprems_of, cskip_flexpairs, cstrip_imp_prems: new cterm operations
 
lcp 
parents: 
668 
diff
changeset
 | 
190  | 
|
| 
2004
 
3411fe560611
New operations on cterms.  Now same names as in Logic
 
paulson 
parents: 
1906 
diff
changeset
 | 
191  | 
(* A1==>...An==>B goes to B, where B is not an implication *)  | 
| 
 
3411fe560611
New operations on cterms.  Now same names as in Logic
 
paulson 
parents: 
1906 
diff
changeset
 | 
192  | 
fun strip_imp_concl ct =  | 
| 8328 | 193  | 
    case term_of ct of (Const("==>", _) $ _ $ _) =>
 | 
| 
10767
 
8fa4aafa7314
Thm: dest_comb, dest_abs, capply, cabs no longer global;
 
wenzelm 
parents: 
10667 
diff
changeset
 | 
194  | 
strip_imp_concl (#2 (Thm.dest_comb ct))  | 
| 
2004
 
3411fe560611
New operations on cterms.  Now same names as in Logic
 
paulson 
parents: 
1906 
diff
changeset
 | 
195  | 
| _ => ct;  | 
| 
 
3411fe560611
New operations on cterms.  Now same names as in Logic
 
paulson 
parents: 
1906 
diff
changeset
 | 
196  | 
|
| 
708
 
8422e50adce0
Pure/drule/cprems_of, cskip_flexpairs, cstrip_imp_prems: new cterm operations
 
lcp 
parents: 
668 
diff
changeset
 | 
197  | 
(*The premises of a theorem, as a cterm list*)  | 
| 
13659
 
3cf622f6b0b2
Removed obsolete functions dealing with flex-flex constraints.
 
berghofe 
parents: 
13650 
diff
changeset
 | 
198  | 
val cprems_of = strip_imp_prems o cprop_of;  | 
| 
708
 
8422e50adce0
Pure/drule/cprems_of, cskip_flexpairs, cstrip_imp_prems: new cterm operations
 
lcp 
parents: 
668 
diff
changeset
 | 
199  | 
|
| 15797 | 200  | 
fun cterm_fun f ct =  | 
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
15949 
diff
changeset
 | 
201  | 
  let val {t, thy, ...} = Thm.rep_cterm ct
 | 
| 
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
15949 
diff
changeset
 | 
202  | 
in Thm.cterm_of thy (f t) end;  | 
| 15797 | 203  | 
|
204  | 
fun ctyp_fun f cT =  | 
|
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
15949 
diff
changeset
 | 
205  | 
  let val {T, thy, ...} = Thm.rep_ctyp cT
 | 
| 
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
15949 
diff
changeset
 | 
206  | 
in Thm.ctyp_of thy (f T) end;  | 
| 15797 | 207  | 
|
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
15949 
diff
changeset
 | 
208  | 
val implies = cterm_of ProtoPure.thy Term.implies;  | 
| 9547 | 209  | 
|
210  | 
(*cterm version of mk_implies*)  | 
|
| 
10767
 
8fa4aafa7314
Thm: dest_comb, dest_abs, capply, cabs no longer global;
 
wenzelm 
parents: 
10667 
diff
changeset
 | 
211  | 
fun mk_implies(A,B) = Thm.capply (Thm.capply implies A) B;  | 
| 9547 | 212  | 
|
213  | 
(*cterm version of list_implies: [A1,...,An], B goes to [|A1;==>;An|]==>B *)  | 
|
214  | 
fun list_implies([], B) = B  | 
|
215  | 
| list_implies(A::AS, B) = mk_implies (A, list_implies(AS,B));  | 
|
216  | 
||
| 15949 | 217  | 
(*cterm version of list_comb: maps (f, [t1,...,tn]) to f(t1,...,tn) *)  | 
218  | 
fun list_comb (f, []) = f  | 
|
219  | 
| list_comb (f, t::ts) = list_comb (Thm.capply f t, ts);  | 
|
220  | 
||
| 
12908
 
53bfe07a7916
New function strip_comb (cterm version of Term.strip_comb).
 
berghofe 
parents: 
12800 
diff
changeset
 | 
221  | 
(*cterm version of strip_comb: maps f(t1,...,tn) to (f, [t1,...,tn]) *)  | 
| 18179 | 222  | 
fun strip_comb ct =  | 
| 
12908
 
53bfe07a7916
New function strip_comb (cterm version of Term.strip_comb).
 
berghofe 
parents: 
12800 
diff
changeset
 | 
223  | 
let  | 
| 
 
53bfe07a7916
New function strip_comb (cterm version of Term.strip_comb).
 
berghofe 
parents: 
12800 
diff
changeset
 | 
224  | 
fun stripc (p as (ct, cts)) =  | 
| 
 
53bfe07a7916
New function strip_comb (cterm version of Term.strip_comb).
 
berghofe 
parents: 
12800 
diff
changeset
 | 
225  | 
let val (ct1, ct2) = Thm.dest_comb ct  | 
| 
 
53bfe07a7916
New function strip_comb (cterm version of Term.strip_comb).
 
berghofe 
parents: 
12800 
diff
changeset
 | 
226  | 
in stripc (ct1, ct2 :: cts) end handle CTERM _ => p  | 
| 
 
53bfe07a7916
New function strip_comb (cterm version of Term.strip_comb).
 
berghofe 
parents: 
12800 
diff
changeset
 | 
227  | 
in stripc (ct, []) end;  | 
| 
 
53bfe07a7916
New function strip_comb (cterm version of Term.strip_comb).
 
berghofe 
parents: 
12800 
diff
changeset
 | 
228  | 
|
| 15262 | 229  | 
(* cterm version of strip_type: maps [T1,...,Tn]--->T to ([T1,T2,...,Tn], T) *)  | 
230  | 
fun strip_type cT = (case Thm.typ_of cT of  | 
|
231  | 
    Type ("fun", _) =>
 | 
|
232  | 
let  | 
|
233  | 
val [cT1, cT2] = Thm.dest_ctyp cT;  | 
|
234  | 
val (cTs, cT') = strip_type cT2  | 
|
235  | 
in (cT1 :: cTs, cT') end  | 
|
236  | 
| _ => ([], cT));  | 
|
237  | 
||
| 15949 | 238  | 
(*Beta-conversion for cterms, where x is an abstraction. Simply returns the rhs  | 
239  | 
of the meta-equality returned by the beta_conversion rule.*)  | 
|
| 18179 | 240  | 
fun beta_conv x y =  | 
| 15949 | 241  | 
#2 (Thm.dest_comb (cprop_of (Thm.beta_conversion false (Thm.capply x y))));  | 
242  | 
||
| 15875 | 243  | 
fun plain_prop_of raw_thm =  | 
244  | 
let  | 
|
245  | 
val thm = Thm.strip_shyps raw_thm;  | 
|
246  | 
    fun err msg = raise THM ("plain_prop_of: " ^ msg, 0, [thm]);
 | 
|
247  | 
    val {hyps, prop, tpairs, ...} = Thm.rep_thm thm;
 | 
|
248  | 
in  | 
|
249  | 
if not (null hyps) then  | 
|
250  | 
err "theorem may not contain hypotheses"  | 
|
251  | 
else if not (null (Thm.extra_shyps thm)) then  | 
|
252  | 
err "theorem may not contain sort hypotheses"  | 
|
253  | 
else if not (null tpairs) then  | 
|
254  | 
err "theorem may not contain flex-flex pairs"  | 
|
255  | 
else prop  | 
|
256  | 
end;  | 
|
257  | 
||
258  | 
||
| 
708
 
8422e50adce0
Pure/drule/cprems_of, cskip_flexpairs, cstrip_imp_prems: new cterm operations
 
lcp 
parents: 
668 
diff
changeset
 | 
259  | 
|
| 
229
 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 
lcp 
parents: 
214 
diff
changeset
 | 
260  | 
(** reading of instantiations **)  | 
| 
 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 
lcp 
parents: 
214 
diff
changeset
 | 
261  | 
|
| 
 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 
lcp 
parents: 
214 
diff
changeset
 | 
262  | 
fun absent ixn =  | 
| 
 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 
lcp 
parents: 
214 
diff
changeset
 | 
263  | 
  error("No such variable in term: " ^ Syntax.string_of_vname ixn);
 | 
| 
 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 
lcp 
parents: 
214 
diff
changeset
 | 
264  | 
|
| 
 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 
lcp 
parents: 
214 
diff
changeset
 | 
265  | 
fun inst_failure ixn =  | 
| 
 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 
lcp 
parents: 
214 
diff
changeset
 | 
266  | 
  error("Instantiation of " ^ Syntax.string_of_vname ixn ^ " fails");
 | 
| 
 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 
lcp 
parents: 
214 
diff
changeset
 | 
267  | 
|
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
15949 
diff
changeset
 | 
268  | 
fun read_insts thy (rtypes,rsorts) (types,sorts) used insts =  | 
| 10403 | 269  | 
let  | 
| 
15442
 
3b75e1b22ff1
Added variants of instantiation functions that operate on pairs of type
 
berghofe 
parents: 
15262 
diff
changeset
 | 
270  | 
fun is_tv ((a, _), _) =  | 
| 
 
3b75e1b22ff1
Added variants of instantiation functions that operate on pairs of type
 
berghofe 
parents: 
15262 
diff
changeset
 | 
271  | 
(case Symbol.explode a of "'" :: _ => true | _ => false);  | 
| 15570 | 272  | 
val (tvs, vs) = List.partition is_tv insts;  | 
| 15797 | 273  | 
fun sort_of ixn = case rsorts ixn of SOME S => S | NONE => absent ixn;  | 
| 
15442
 
3b75e1b22ff1
Added variants of instantiation functions that operate on pairs of type
 
berghofe 
parents: 
15262 
diff
changeset
 | 
274  | 
fun readT (ixn, st) =  | 
| 15797 | 275  | 
let val S = sort_of ixn;  | 
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
15949 
diff
changeset
 | 
276  | 
val T = Sign.read_typ (thy,sorts) st;  | 
| 
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
15949 
diff
changeset
 | 
277  | 
in if Sign.typ_instance thy (T, TVar(ixn,S)) then (ixn,T)  | 
| 
4281
 
6c6073b13600
Added read_def_cterms for simultaneous reading/typing of terms under
 
nipkow 
parents: 
4270 
diff
changeset
 | 
278  | 
else inst_failure ixn  | 
| 
 
6c6073b13600
Added read_def_cterms for simultaneous reading/typing of terms under
 
nipkow 
parents: 
4270 
diff
changeset
 | 
279  | 
end  | 
| 
 
6c6073b13600
Added read_def_cterms for simultaneous reading/typing of terms under
 
nipkow 
parents: 
4270 
diff
changeset
 | 
280  | 
val tye = map readT tvs;  | 
| 
 
6c6073b13600
Added read_def_cterms for simultaneous reading/typing of terms under
 
nipkow 
parents: 
4270 
diff
changeset
 | 
281  | 
fun mkty(ixn,st) = (case rtypes ixn of  | 
| 15531 | 282  | 
SOME T => (ixn,(st,typ_subst_TVars tye T))  | 
283  | 
| NONE => absent ixn);  | 
|
| 
4281
 
6c6073b13600
Added read_def_cterms for simultaneous reading/typing of terms under
 
nipkow 
parents: 
4270 
diff
changeset
 | 
284  | 
val ixnsTs = map mkty vs;  | 
| 
 
6c6073b13600
Added read_def_cterms for simultaneous reading/typing of terms under
 
nipkow 
parents: 
4270 
diff
changeset
 | 
285  | 
val ixns = map fst ixnsTs  | 
| 
 
6c6073b13600
Added read_def_cterms for simultaneous reading/typing of terms under
 
nipkow 
parents: 
4270 
diff
changeset
 | 
286  | 
and sTs = map snd ixnsTs  | 
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
15949 
diff
changeset
 | 
287  | 
val (cts,tye2) = read_def_cterms(thy,types,sorts) used false sTs;  | 
| 
4281
 
6c6073b13600
Added read_def_cterms for simultaneous reading/typing of terms under
 
nipkow 
parents: 
4270 
diff
changeset
 | 
288  | 
fun mkcVar(ixn,T) =  | 
| 
 
6c6073b13600
Added read_def_cterms for simultaneous reading/typing of terms under
 
nipkow 
parents: 
4270 
diff
changeset
 | 
289  | 
let val U = typ_subst_TVars tye2 T  | 
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
15949 
diff
changeset
 | 
290  | 
in cterm_of thy (Var(ixn,U)) end  | 
| 
4281
 
6c6073b13600
Added read_def_cterms for simultaneous reading/typing of terms under
 
nipkow 
parents: 
4270 
diff
changeset
 | 
291  | 
val ixnTs = ListPair.zip(ixns, map snd sTs)  | 
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
15949 
diff
changeset
 | 
292  | 
in (map (fn (ixn, T) => (ctyp_of thy (TVar (ixn, sort_of ixn)),  | 
| 
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
15949 
diff
changeset
 | 
293  | 
ctyp_of thy T)) (tye2 @ tye),  | 
| 
4281
 
6c6073b13600
Added read_def_cterms for simultaneous reading/typing of terms under
 
nipkow 
parents: 
4270 
diff
changeset
 | 
294  | 
ListPair.zip(map mkcVar ixnTs,cts))  | 
| 
 
6c6073b13600
Added read_def_cterms for simultaneous reading/typing of terms under
 
nipkow 
parents: 
4270 
diff
changeset
 | 
295  | 
end;  | 
| 
229
 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 
lcp 
parents: 
214 
diff
changeset
 | 
296  | 
|
| 
 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 
lcp 
parents: 
214 
diff
changeset
 | 
297  | 
|
| 
252
 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 
wenzelm 
parents: 
229 
diff
changeset
 | 
298  | 
(*** Find the type (sort) associated with a (T)Var or (T)Free in a term  | 
| 0 | 299  | 
Used for establishing default types (of variables) and sorts (of  | 
300  | 
type variables) when reading another term.  | 
|
301  | 
Index -1 indicates that a (T)Free rather than a (T)Var is wanted.  | 
|
302  | 
***)  | 
|
303  | 
||
304  | 
fun types_sorts thm =  | 
|
| 15669 | 305  | 
    let val {prop, hyps, tpairs, ...} = Thm.rep_thm thm;
 | 
306  | 
(* bogus term! *)  | 
|
| 18179 | 307  | 
val big = Term.list_comb  | 
| 15949 | 308  | 
(Term.list_comb (prop, hyps), Thm.terms_of_tpairs tpairs);  | 
| 
252
 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 
wenzelm 
parents: 
229 
diff
changeset
 | 
309  | 
val vars = map dest_Var (term_vars big);  | 
| 
 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 
wenzelm 
parents: 
229 
diff
changeset
 | 
310  | 
val frees = map dest_Free (term_frees big);  | 
| 
 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 
wenzelm 
parents: 
229 
diff
changeset
 | 
311  | 
val tvars = term_tvars big;  | 
| 
 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 
wenzelm 
parents: 
229 
diff
changeset
 | 
312  | 
val tfrees = term_tfrees big;  | 
| 17325 | 313  | 
fun typ(a,i) = if i<0 then AList.lookup (op =) frees a else AList.lookup (op =) vars (a,i);  | 
314  | 
fun sort(a,i) = if i<0 then AList.lookup (op =) tfrees a else AList.lookup (op =) tvars (a,i);  | 
|
| 0 | 315  | 
in (typ,sort) end;  | 
316  | 
||
| 15669 | 317  | 
fun add_used thm used =  | 
318  | 
  let val {prop, hyps, tpairs, ...} = Thm.rep_thm thm in
 | 
|
319  | 
add_term_tvarnames (prop, used)  | 
|
320  | 
|> fold (curry add_term_tvarnames) hyps  | 
|
321  | 
|> fold (curry add_term_tvarnames) (Thm.terms_of_tpairs tpairs)  | 
|
322  | 
end;  | 
|
323  | 
||
| 7636 | 324  | 
|
| 9455 | 325  | 
|
326  | 
(** basic attributes **)  | 
|
327  | 
||
328  | 
(* dependent rules *)  | 
|
329  | 
||
330  | 
fun rule_attribute f (x, thm) = (x, (f x thm));  | 
|
331  | 
||
332  | 
||
333  | 
(* add / delete tags *)  | 
|
334  | 
||
335  | 
fun map_tags f thm =  | 
|
336  | 
Thm.put_name_tags (Thm.name_of_thm thm, f (#2 (Thm.get_name_tags thm))) thm;  | 
|
337  | 
||
338  | 
fun tag_rule tg = map_tags (fn tgs => if tg mem tgs then tgs else tgs @ [tg]);  | 
|
339  | 
fun untag_rule s = map_tags (filter_out (equal s o #1));  | 
|
340  | 
||
341  | 
fun tag tg x = rule_attribute (K (tag_rule tg)) x;  | 
|
342  | 
fun untag s x = rule_attribute (K (untag_rule s)) x;  | 
|
343  | 
||
344  | 
fun simple_tag name x = tag (name, []) x;  | 
|
345  | 
||
| 11741 | 346  | 
|
347  | 
(* theorem kinds *)  | 
|
348  | 
||
349  | 
val theoremK = "theorem";  | 
|
350  | 
val lemmaK = "lemma";  | 
|
351  | 
val corollaryK = "corollary";  | 
|
352  | 
val internalK = "internal";  | 
|
| 9455 | 353  | 
|
| 11741 | 354  | 
fun get_kind thm =  | 
| 17325 | 355  | 
(case AList.lookup (op =) ((#2 o Thm.get_name_tags) thm) "kind" of  | 
| 15531 | 356  | 
SOME (k :: _) => k  | 
| 11741 | 357  | 
| _ => "unknown");  | 
358  | 
||
359  | 
fun kind_rule k = tag_rule ("kind", [k]) o untag_rule "kind";
 | 
|
| 12710 | 360  | 
fun kind k x = if k = "" then x else rule_attribute (K (kind_rule k)) x;  | 
| 11741 | 361  | 
fun kind_internal x = kind internalK x;  | 
| 18468 | 362  | 
fun has_internal tags = exists (fn ("kind", [k]) => k = internalK | _ => false) tags;
 | 
363  | 
val is_internal = has_internal o Thm.tags_of_thm;  | 
|
| 9455 | 364  | 
|
365  | 
||
366  | 
||
| 0 | 367  | 
(** Standardization of rules **)  | 
368  | 
||
| 18025 | 369  | 
(*vars in left-to-right order*)  | 
370  | 
fun tvars_of_terms ts = rev (fold Term.add_tvars ts []);  | 
|
371  | 
fun vars_of_terms ts = rev (fold Term.add_vars ts []);  | 
|
372  | 
fun tvars_of thm = tvars_of_terms [Thm.full_prop_of thm];  | 
|
373  | 
fun vars_of thm = vars_of_terms [Thm.full_prop_of thm];  | 
|
374  | 
||
| 18129 | 375  | 
fun fold_terms f th =  | 
376  | 
  let val {hyps, tpairs, prop, ...} = Thm.rep_thm th
 | 
|
377  | 
in f prop #> fold (fn (t, u) => f t #> f u) tpairs #> fold f hyps end;  | 
|
378  | 
||
379  | 
fun tfrees_of th = rev (fold_terms Term.add_tfrees th []);  | 
|
380  | 
fun frees_of th = rev (fold_terms Term.add_frees th []);  | 
|
381  | 
||
| 7636 | 382  | 
(*Strip extraneous shyps as far as possible*)  | 
383  | 
fun strip_shyps_warning thm =  | 
|
384  | 
let  | 
|
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
15949 
diff
changeset
 | 
385  | 
val str_of_sort = Pretty.str_of o Sign.pretty_sort (Thm.theory_of_thm thm);  | 
| 7636 | 386  | 
val thm' = Thm.strip_shyps thm;  | 
387  | 
val xshyps = Thm.extra_shyps thm';  | 
|
388  | 
in  | 
|
389  | 
if null xshyps then ()  | 
|
390  | 
    else warning ("Pending sort hypotheses: " ^ commas (map str_of_sort xshyps));
 | 
|
391  | 
thm'  | 
|
392  | 
end;  | 
|
393  | 
||
| 0 | 394  | 
(*Generalization over a list of variables, IGNORING bad ones*)  | 
395  | 
fun forall_intr_list [] th = th  | 
|
396  | 
| forall_intr_list (y::ys) th =  | 
|
| 
252
 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 
wenzelm 
parents: 
229 
diff
changeset
 | 
397  | 
let val gth = forall_intr_list ys th  | 
| 
 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 
wenzelm 
parents: 
229 
diff
changeset
 | 
398  | 
in forall_intr y gth handle THM _ => gth end;  | 
| 0 | 399  | 
|
400  | 
(*Generalization over all suitable Free variables*)  | 
|
401  | 
fun forall_intr_frees th =  | 
|
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
15949 
diff
changeset
 | 
402  | 
    let val {prop,thy,...} = rep_thm th
 | 
| 0 | 403  | 
in forall_intr_list  | 
| 16983 | 404  | 
(map (cterm_of thy) (sort Term.term_ord (term_frees prop)))  | 
| 0 | 405  | 
th  | 
406  | 
end;  | 
|
407  | 
||
| 7898 | 408  | 
val forall_elim_var = PureThy.forall_elim_var;  | 
409  | 
val forall_elim_vars = PureThy.forall_elim_vars;  | 
|
| 0 | 410  | 
|
| 18025 | 411  | 
fun outer_params t =  | 
412  | 
let  | 
|
413  | 
val vs = Term.strip_all_vars t;  | 
|
| 18375 | 414  | 
val xs = Term.variantlist (map (perhaps (try Syntax.dest_skolem) o #1) vs, []);  | 
| 18025 | 415  | 
in xs ~~ map #2 vs end;  | 
416  | 
||
417  | 
(*generalize outermost parameters*)  | 
|
418  | 
fun gen_all th =  | 
|
| 
12719
 
41e0d086f8b6
improved forall_elim_vars_safe (no longer invents new indexes);
 
wenzelm 
parents: 
12710 
diff
changeset
 | 
419  | 
let  | 
| 18025 | 420  | 
    val {thy, prop, maxidx, ...} = Thm.rep_thm th;
 | 
421  | 
val cert = Thm.cterm_of thy;  | 
|
422  | 
fun elim (x, T) = Thm.forall_elim (cert (Var ((x, maxidx + 1), T)));  | 
|
423  | 
in fold elim (outer_params prop) th end;  | 
|
424  | 
||
425  | 
(*lift vars wrt. outermost goal parameters  | 
|
| 18118 | 426  | 
-- reverses the effect of gen_all modulo higher-order unification*)  | 
| 18025 | 427  | 
fun lift_all goal th =  | 
428  | 
let  | 
|
429  | 
val thy = Theory.merge (Thm.theory_of_cterm goal, Thm.theory_of_thm th);  | 
|
430  | 
val cert = Thm.cterm_of thy;  | 
|
431  | 
    val {maxidx, ...} = Thm.rep_thm th;
 | 
|
432  | 
val ps = outer_params (Thm.term_of goal)  | 
|
433  | 
|> map (fn (x, T) => Var ((x, maxidx + 1), Logic.incr_tvar (maxidx + 1) T));  | 
|
434  | 
val Ts = map Term.fastype_of ps;  | 
|
435  | 
val inst = vars_of th |> map (fn (xi, T) =>  | 
|
436  | 
(cert (Var (xi, T)), cert (Term.list_comb (Var (xi, Ts ---> T), ps))));  | 
|
437  | 
in  | 
|
438  | 
th |> Thm.instantiate ([], inst)  | 
|
439  | 
|> fold_rev (Thm.forall_intr o cert) ps  | 
|
440  | 
end;  | 
|
441  | 
||
| 9554 | 442  | 
|
| 
16949
 
ea65d75e0ce1
tuned gen_all, forall_elim_list, implies_intr_list, standard;
 
wenzelm 
parents: 
16861 
diff
changeset
 | 
443  | 
(*specialization over a list of cterms*)  | 
| 
 
ea65d75e0ce1
tuned gen_all, forall_elim_list, implies_intr_list, standard;
 
wenzelm 
parents: 
16861 
diff
changeset
 | 
444  | 
val forall_elim_list = fold forall_elim;  | 
| 0 | 445  | 
|
| 
16949
 
ea65d75e0ce1
tuned gen_all, forall_elim_list, implies_intr_list, standard;
 
wenzelm 
parents: 
16861 
diff
changeset
 | 
446  | 
(*maps A1,...,An |- B to [| A1;...;An |] ==> B*)  | 
| 
 
ea65d75e0ce1
tuned gen_all, forall_elim_list, implies_intr_list, standard;
 
wenzelm 
parents: 
16861 
diff
changeset
 | 
447  | 
val implies_intr_list = fold_rev implies_intr;  | 
| 0 | 448  | 
|
| 
16949
 
ea65d75e0ce1
tuned gen_all, forall_elim_list, implies_intr_list, standard;
 
wenzelm 
parents: 
16861 
diff
changeset
 | 
449  | 
(*maps [| A1;...;An |] ==> B and [A1,...,An] to B*)  | 
| 15570 | 450  | 
fun implies_elim_list impth ths = Library.foldl (uncurry implies_elim) (impth,ths);  | 
| 0 | 451  | 
|
452  | 
(*Reset Var indexes to zero, renaming to preserve distinctness*)  | 
|
| 
252
 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 
wenzelm 
parents: 
229 
diff
changeset
 | 
453  | 
fun zero_var_indexes th =  | 
| 
16949
 
ea65d75e0ce1
tuned gen_all, forall_elim_list, implies_intr_list, standard;
 
wenzelm 
parents: 
16861 
diff
changeset
 | 
454  | 
let  | 
| 
 
ea65d75e0ce1
tuned gen_all, forall_elim_list, implies_intr_list, standard;
 
wenzelm 
parents: 
16861 
diff
changeset
 | 
455  | 
val thy = Thm.theory_of_thm th;  | 
| 
 
ea65d75e0ce1
tuned gen_all, forall_elim_list, implies_intr_list, standard;
 
wenzelm 
parents: 
16861 
diff
changeset
 | 
456  | 
val certT = Thm.ctyp_of thy and cert = Thm.cterm_of thy;  | 
| 
 
ea65d75e0ce1
tuned gen_all, forall_elim_list, implies_intr_list, standard;
 
wenzelm 
parents: 
16861 
diff
changeset
 | 
457  | 
val (instT, inst) = Term.zero_var_indexes_inst (Thm.full_prop_of th);  | 
| 
 
ea65d75e0ce1
tuned gen_all, forall_elim_list, implies_intr_list, standard;
 
wenzelm 
parents: 
16861 
diff
changeset
 | 
458  | 
val cinstT = map (fn (v, T) => (certT (TVar v), certT T)) instT;  | 
| 
 
ea65d75e0ce1
tuned gen_all, forall_elim_list, implies_intr_list, standard;
 
wenzelm 
parents: 
16861 
diff
changeset
 | 
459  | 
val cinst = map (fn (v, t) => (cert (Var v), cert t)) inst;  | 
| 
 
ea65d75e0ce1
tuned gen_all, forall_elim_list, implies_intr_list, standard;
 
wenzelm 
parents: 
16861 
diff
changeset
 | 
460  | 
in Thm.adjust_maxidx_thm (Thm.instantiate (cinstT, cinst) th) end;  | 
| 0 | 461  | 
|
462  | 
||
| 14394 | 463  | 
(** Standard form of object-rule: no hypotheses, flexflex constraints,  | 
464  | 
Frees, or outer quantifiers; all generality expressed by Vars of index 0.**)  | 
|
| 10515 | 465  | 
|
| 16595 | 466  | 
(*Discharge all hypotheses.*)  | 
467  | 
fun implies_intr_hyps th =  | 
|
468  | 
fold Thm.implies_intr (#hyps (Thm.crep_thm th)) th;  | 
|
469  | 
||
| 14394 | 470  | 
(*Squash a theorem's flexflex constraints provided it can be done uniquely.  | 
471  | 
This step can lose information.*)  | 
|
| 
14387
 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 
paulson 
parents: 
14340 
diff
changeset
 | 
472  | 
fun flexflex_unique th =  | 
| 17713 | 473  | 
if null (tpairs_of th) then th else  | 
| 
14387
 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 
paulson 
parents: 
14340 
diff
changeset
 | 
474  | 
case Seq.chop (2, flexflex_rule th) of  | 
| 
 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 
paulson 
parents: 
14340 
diff
changeset
 | 
475  | 
([th],_) => th  | 
| 
 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 
paulson 
parents: 
14340 
diff
changeset
 | 
476  | 
    | ([],_)   => raise THM("flexflex_unique: impossible constraints", 0, [th])
 | 
| 
 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 
paulson 
parents: 
14340 
diff
changeset
 | 
477  | 
    |      _   => raise THM("flexflex_unique: multiple unifiers", 0, [th]);
 | 
| 
 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 
paulson 
parents: 
14340 
diff
changeset
 | 
478  | 
|
| 10515 | 479  | 
fun close_derivation thm =  | 
480  | 
  if Thm.get_name_tags thm = ("", []) then Thm.name_thm ("", thm)
 | 
|
481  | 
else thm;  | 
|
482  | 
||
| 
16949
 
ea65d75e0ce1
tuned gen_all, forall_elim_list, implies_intr_list, standard;
 
wenzelm 
parents: 
16861 
diff
changeset
 | 
483  | 
val standard' =  | 
| 
 
ea65d75e0ce1
tuned gen_all, forall_elim_list, implies_intr_list, standard;
 
wenzelm 
parents: 
16861 
diff
changeset
 | 
484  | 
implies_intr_hyps  | 
| 
 
ea65d75e0ce1
tuned gen_all, forall_elim_list, implies_intr_list, standard;
 
wenzelm 
parents: 
16861 
diff
changeset
 | 
485  | 
#> forall_intr_frees  | 
| 
 
ea65d75e0ce1
tuned gen_all, forall_elim_list, implies_intr_list, standard;
 
wenzelm 
parents: 
16861 
diff
changeset
 | 
486  | 
#> `(#maxidx o Thm.rep_thm)  | 
| 
 
ea65d75e0ce1
tuned gen_all, forall_elim_list, implies_intr_list, standard;
 
wenzelm 
parents: 
16861 
diff
changeset
 | 
487  | 
#-> (fn maxidx =>  | 
| 
 
ea65d75e0ce1
tuned gen_all, forall_elim_list, implies_intr_list, standard;
 
wenzelm 
parents: 
16861 
diff
changeset
 | 
488  | 
forall_elim_vars (maxidx + 1)  | 
| 
 
ea65d75e0ce1
tuned gen_all, forall_elim_list, implies_intr_list, standard;
 
wenzelm 
parents: 
16861 
diff
changeset
 | 
489  | 
#> strip_shyps_warning  | 
| 
 
ea65d75e0ce1
tuned gen_all, forall_elim_list, implies_intr_list, standard;
 
wenzelm 
parents: 
16861 
diff
changeset
 | 
490  | 
#> zero_var_indexes  | 
| 
 
ea65d75e0ce1
tuned gen_all, forall_elim_list, implies_intr_list, standard;
 
wenzelm 
parents: 
16861 
diff
changeset
 | 
491  | 
#> Thm.varifyT  | 
| 
 
ea65d75e0ce1
tuned gen_all, forall_elim_list, implies_intr_list, standard;
 
wenzelm 
parents: 
16861 
diff
changeset
 | 
492  | 
#> Thm.compress);  | 
| 
1218
 
59ed8ef1a3a1
modified pretty_thm, standard, eq_thm to handle shyps;
 
wenzelm 
parents: 
1194 
diff
changeset
 | 
493  | 
|
| 
16949
 
ea65d75e0ce1
tuned gen_all, forall_elim_list, implies_intr_list, standard;
 
wenzelm 
parents: 
16861 
diff
changeset
 | 
494  | 
val standard =  | 
| 
 
ea65d75e0ce1
tuned gen_all, forall_elim_list, implies_intr_list, standard;
 
wenzelm 
parents: 
16861 
diff
changeset
 | 
495  | 
flexflex_unique  | 
| 
 
ea65d75e0ce1
tuned gen_all, forall_elim_list, implies_intr_list, standard;
 
wenzelm 
parents: 
16861 
diff
changeset
 | 
496  | 
#> standard'  | 
| 
 
ea65d75e0ce1
tuned gen_all, forall_elim_list, implies_intr_list, standard;
 
wenzelm 
parents: 
16861 
diff
changeset
 | 
497  | 
#> close_derivation;  | 
| 
11512
 
da3a96ab5630
Some basic rules are now stored with "open" derivations, to facilitate
 
berghofe 
parents: 
11163 
diff
changeset
 | 
498  | 
|
| 
16949
 
ea65d75e0ce1
tuned gen_all, forall_elim_list, implies_intr_list, standard;
 
wenzelm 
parents: 
16861 
diff
changeset
 | 
499  | 
val local_standard =  | 
| 
 
ea65d75e0ce1
tuned gen_all, forall_elim_list, implies_intr_list, standard;
 
wenzelm 
parents: 
16861 
diff
changeset
 | 
500  | 
strip_shyps  | 
| 
 
ea65d75e0ce1
tuned gen_all, forall_elim_list, implies_intr_list, standard;
 
wenzelm 
parents: 
16861 
diff
changeset
 | 
501  | 
#> zero_var_indexes  | 
| 
 
ea65d75e0ce1
tuned gen_all, forall_elim_list, implies_intr_list, standard;
 
wenzelm 
parents: 
16861 
diff
changeset
 | 
502  | 
#> Thm.compress  | 
| 
 
ea65d75e0ce1
tuned gen_all, forall_elim_list, implies_intr_list, standard;
 
wenzelm 
parents: 
16861 
diff
changeset
 | 
503  | 
#> close_derivation;  | 
| 12005 | 504  | 
|
| 0 | 505  | 
|
| 8328 | 506  | 
(*Convert all Vars in a theorem to Frees. Also return a function for  | 
| 
4610
 
b1322be47244
Tidying; rotate_prems; moved freeze_thaw from tactic.ML
 
paulson 
parents: 
4588 
diff
changeset
 | 
507  | 
reversing that operation. DOES NOT WORK FOR TYPE VARIABLES.  | 
| 
 
b1322be47244
Tidying; rotate_prems; moved freeze_thaw from tactic.ML
 
paulson 
parents: 
4588 
diff
changeset
 | 
508  | 
Similar code in type/freeze_thaw*)  | 
| 
15495
 
50fde6f04e4c
new treatment of demodulation in proof reconstruction
 
paulson 
parents: 
15442 
diff
changeset
 | 
509  | 
|
| 
 
50fde6f04e4c
new treatment of demodulation in proof reconstruction
 
paulson 
parents: 
15442 
diff
changeset
 | 
510  | 
fun freeze_thaw_robust th =  | 
| 
 
50fde6f04e4c
new treatment of demodulation in proof reconstruction
 
paulson 
parents: 
15442 
diff
changeset
 | 
511  | 
let val fth = freezeT th  | 
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
15949 
diff
changeset
 | 
512  | 
     val {prop, tpairs, thy, ...} = rep_thm fth
 | 
| 
15495
 
50fde6f04e4c
new treatment of demodulation in proof reconstruction
 
paulson 
parents: 
15442 
diff
changeset
 | 
513  | 
in  | 
| 
15574
 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 
skalberg 
parents: 
15570 
diff
changeset
 | 
514  | 
case foldr add_term_vars [] (prop :: Thm.terms_of_tpairs tpairs) of  | 
| 
15495
 
50fde6f04e4c
new treatment of demodulation in proof reconstruction
 
paulson 
parents: 
15442 
diff
changeset
 | 
515  | 
[] => (fth, fn i => fn x => x) (*No vars: nothing to do!*)  | 
| 
 
50fde6f04e4c
new treatment of demodulation in proof reconstruction
 
paulson 
parents: 
15442 
diff
changeset
 | 
516  | 
| vars =>  | 
| 
 
50fde6f04e4c
new treatment of demodulation in proof reconstruction
 
paulson 
parents: 
15442 
diff
changeset
 | 
517  | 
let fun newName (Var(ix,_), pairs) =  | 
| 
 
50fde6f04e4c
new treatment of demodulation in proof reconstruction
 
paulson 
parents: 
15442 
diff
changeset
 | 
518  | 
let val v = gensym (string_of_indexname ix)  | 
| 
 
50fde6f04e4c
new treatment of demodulation in proof reconstruction
 
paulson 
parents: 
15442 
diff
changeset
 | 
519  | 
in ((ix,v)::pairs) end;  | 
| 
15574
 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 
skalberg 
parents: 
15570 
diff
changeset
 | 
520  | 
val alist = foldr newName [] vars  | 
| 
15495
 
50fde6f04e4c
new treatment of demodulation in proof reconstruction
 
paulson 
parents: 
15442 
diff
changeset
 | 
521  | 
fun mk_inst (Var(v,T)) =  | 
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
15949 
diff
changeset
 | 
522  | 
(cterm_of thy (Var(v,T)),  | 
| 17325 | 523  | 
cterm_of thy (Free(((the o AList.lookup (op =) alist) v), T)))  | 
| 
15495
 
50fde6f04e4c
new treatment of demodulation in proof reconstruction
 
paulson 
parents: 
15442 
diff
changeset
 | 
524  | 
val insts = map mk_inst vars  | 
| 
 
50fde6f04e4c
new treatment of demodulation in proof reconstruction
 
paulson 
parents: 
15442 
diff
changeset
 | 
525  | 
fun thaw i th' = (*i is non-negative increment for Var indexes*)  | 
| 
 
50fde6f04e4c
new treatment of demodulation in proof reconstruction
 
paulson 
parents: 
15442 
diff
changeset
 | 
526  | 
th' |> forall_intr_list (map #2 insts)  | 
| 
 
50fde6f04e4c
new treatment of demodulation in proof reconstruction
 
paulson 
parents: 
15442 
diff
changeset
 | 
527  | 
|> forall_elim_list (map (Thm.cterm_incr_indexes i o #1) insts)  | 
| 
 
50fde6f04e4c
new treatment of demodulation in proof reconstruction
 
paulson 
parents: 
15442 
diff
changeset
 | 
528  | 
in (Thm.instantiate ([],insts) fth, thaw) end  | 
| 
 
50fde6f04e4c
new treatment of demodulation in proof reconstruction
 
paulson 
parents: 
15442 
diff
changeset
 | 
529  | 
end;  | 
| 
 
50fde6f04e4c
new treatment of demodulation in proof reconstruction
 
paulson 
parents: 
15442 
diff
changeset
 | 
530  | 
|
| 
 
50fde6f04e4c
new treatment of demodulation in proof reconstruction
 
paulson 
parents: 
15442 
diff
changeset
 | 
531  | 
(*Basic version of the function above. No option to rename Vars apart in thaw.  | 
| 
 
50fde6f04e4c
new treatment of demodulation in proof reconstruction
 
paulson 
parents: 
15442 
diff
changeset
 | 
532  | 
The Frees created from Vars have nice names.*)  | 
| 
4610
 
b1322be47244
Tidying; rotate_prems; moved freeze_thaw from tactic.ML
 
paulson 
parents: 
4588 
diff
changeset
 | 
533  | 
fun freeze_thaw th =  | 
| 
7248
 
322151fe6f02
new primitive rule permute_prems to underlie defer_tac and rotate_prems
 
paulson 
parents: 
6995 
diff
changeset
 | 
534  | 
let val fth = freezeT th  | 
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
15949 
diff
changeset
 | 
535  | 
     val {prop, tpairs, thy, ...} = rep_thm fth
 | 
| 
7248
 
322151fe6f02
new primitive rule permute_prems to underlie defer_tac and rotate_prems
 
paulson 
parents: 
6995 
diff
changeset
 | 
536  | 
in  | 
| 
15574
 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 
skalberg 
parents: 
15570 
diff
changeset
 | 
537  | 
case foldr add_term_vars [] (prop :: Thm.terms_of_tpairs tpairs) of  | 
| 
7248
 
322151fe6f02
new primitive rule permute_prems to underlie defer_tac and rotate_prems
 
paulson 
parents: 
6995 
diff
changeset
 | 
538  | 
[] => (fth, fn x => x)  | 
| 
 
322151fe6f02
new primitive rule permute_prems to underlie defer_tac and rotate_prems
 
paulson 
parents: 
6995 
diff
changeset
 | 
539  | 
| vars =>  | 
| 8328 | 540  | 
let fun newName (Var(ix,_), (pairs,used)) =  | 
541  | 
let val v = variant used (string_of_indexname ix)  | 
|
542  | 
in ((ix,v)::pairs, v::used) end;  | 
|
| 
15574
 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 
skalberg 
parents: 
15570 
diff
changeset
 | 
543  | 
val (alist, _) = foldr newName ([], Library.foldr add_term_names  | 
| 
 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 
skalberg 
parents: 
15570 
diff
changeset
 | 
544  | 
(prop :: Thm.terms_of_tpairs tpairs, [])) vars  | 
| 8328 | 545  | 
fun mk_inst (Var(v,T)) =  | 
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
15949 
diff
changeset
 | 
546  | 
(cterm_of thy (Var(v,T)),  | 
| 17325 | 547  | 
cterm_of thy (Free(((the o AList.lookup (op =) alist) v), T)))  | 
| 8328 | 548  | 
val insts = map mk_inst vars  | 
549  | 
fun thaw th' =  | 
|
550  | 
th' |> forall_intr_list (map #2 insts)  | 
|
551  | 
|> forall_elim_list (map #1 insts)  | 
|
552  | 
in (Thm.instantiate ([],insts) fth, thaw) end  | 
|
| 
7248
 
322151fe6f02
new primitive rule permute_prems to underlie defer_tac and rotate_prems
 
paulson 
parents: 
6995 
diff
changeset
 | 
553  | 
end;  | 
| 
4610
 
b1322be47244
Tidying; rotate_prems; moved freeze_thaw from tactic.ML
 
paulson 
parents: 
4588 
diff
changeset
 | 
554  | 
|
| 
7248
 
322151fe6f02
new primitive rule permute_prems to underlie defer_tac and rotate_prems
 
paulson 
parents: 
6995 
diff
changeset
 | 
555  | 
(*Rotates a rule's premises to the left by k*)  | 
| 
 
322151fe6f02
new primitive rule permute_prems to underlie defer_tac and rotate_prems
 
paulson 
parents: 
6995 
diff
changeset
 | 
556  | 
val rotate_prems = permute_prems 0;  | 
| 
4610
 
b1322be47244
Tidying; rotate_prems; moved freeze_thaw from tactic.ML
 
paulson 
parents: 
4588 
diff
changeset
 | 
557  | 
|
| 11163 | 558  | 
(* permute prems, where the i-th position in the argument list (counting from 0)  | 
559  | 
gives the position within the original thm to be transferred to position i.  | 
|
560  | 
Any remaining trailing positions are left unchanged. *)  | 
|
561  | 
val rearrange_prems = let  | 
|
562  | 
fun rearr new [] thm = thm  | 
|
| 11815 | 563  | 
| rearr new (p::ps) thm = rearr (new+1)  | 
| 11163 | 564  | 
(map (fn q => if new<=q andalso q<p then q+1 else q) ps)  | 
565  | 
(permute_prems (new+1) (new-p) (permute_prems new (p-new) thm))  | 
|
566  | 
in rearr 0 end;  | 
|
| 
4610
 
b1322be47244
Tidying; rotate_prems; moved freeze_thaw from tactic.ML
 
paulson 
parents: 
4588 
diff
changeset
 | 
567  | 
|
| 
252
 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 
wenzelm 
parents: 
229 
diff
changeset
 | 
568  | 
(*Assume a new formula, read following the same conventions as axioms.  | 
| 0 | 569  | 
Generalizes over Free variables,  | 
570  | 
creates the assumption, and then strips quantifiers.  | 
|
571  | 
Example is [| ALL x:?A. ?P(x) |] ==> [| ?P(?a) |]  | 
|
| 
252
 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 
wenzelm 
parents: 
229 
diff
changeset
 | 
572  | 
[ !(A,P,a)[| ALL x:A. P(x) |] ==> [| P(a) |] ] *)  | 
| 0 | 573  | 
fun assume_ax thy sP =  | 
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
15949 
diff
changeset
 | 
574  | 
let val prop = Logic.close_form (term_of (read_cterm thy (sP, propT)))  | 
| 
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
15949 
diff
changeset
 | 
575  | 
in forall_elim_vars 0 (Thm.assume (cterm_of thy prop)) end;  | 
| 0 | 576  | 
|
| 
252
 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 
wenzelm 
parents: 
229 
diff
changeset
 | 
577  | 
(*Resolution: exactly one resolvent must be produced.*)  | 
| 0 | 578  | 
fun tha RSN (i,thb) =  | 
| 4270 | 579  | 
case Seq.chop (2, biresolution false [(false,tha)] i thb) of  | 
| 0 | 580  | 
([th],_) => th  | 
581  | 
    | ([],_)   => raise THM("RSN: no unifiers", i, [tha,thb])
 | 
|
582  | 
    |      _   => raise THM("RSN: multiple unifiers", i, [tha,thb]);
 | 
|
583  | 
||
584  | 
(*resolution: P==>Q, Q==>R gives P==>R. *)  | 
|
585  | 
fun tha RS thb = tha RSN (1,thb);  | 
|
586  | 
||
587  | 
(*For joining lists of rules*)  | 
|
| 
252
 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 
wenzelm 
parents: 
229 
diff
changeset
 | 
588  | 
fun thas RLN (i,thbs) =  | 
| 0 | 589  | 
let val resolve = biresolution false (map (pair false) thas) i  | 
| 4270 | 590  | 
fun resb thb = Seq.list_of (resolve thb) handle THM _ => []  | 
| 
2672
 
85d7e800d754
Replaced "flat" by the Basis Library function List.concat
 
paulson 
parents: 
2266 
diff
changeset
 | 
591  | 
in List.concat (map resb thbs) end;  | 
| 0 | 592  | 
|
593  | 
fun thas RL thbs = thas RLN (1,thbs);  | 
|
594  | 
||
| 
11
 
d0e17c42dbb4
Added MRS, MRL from ZF/ROOT.ML.  These support forward proof, resolving a
 
lcp 
parents: 
0 
diff
changeset
 | 
595  | 
(*Resolve a list of rules against bottom_rl from right to left;  | 
| 
 
d0e17c42dbb4
Added MRS, MRL from ZF/ROOT.ML.  These support forward proof, resolving a
 
lcp 
parents: 
0 
diff
changeset
 | 
596  | 
makes proof trees*)  | 
| 
252
 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 
wenzelm 
parents: 
229 
diff
changeset
 | 
597  | 
fun rls MRS bottom_rl =  | 
| 
11
 
d0e17c42dbb4
Added MRS, MRL from ZF/ROOT.ML.  These support forward proof, resolving a
 
lcp 
parents: 
0 
diff
changeset
 | 
598  | 
let fun rs_aux i [] = bottom_rl  | 
| 
252
 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 
wenzelm 
parents: 
229 
diff
changeset
 | 
599  | 
| rs_aux i (rl::rls) = rl RSN (i, rs_aux (i+1) rls)  | 
| 
11
 
d0e17c42dbb4
Added MRS, MRL from ZF/ROOT.ML.  These support forward proof, resolving a
 
lcp 
parents: 
0 
diff
changeset
 | 
600  | 
in rs_aux 1 rls end;  | 
| 
 
d0e17c42dbb4
Added MRS, MRL from ZF/ROOT.ML.  These support forward proof, resolving a
 
lcp 
parents: 
0 
diff
changeset
 | 
601  | 
|
| 
 
d0e17c42dbb4
Added MRS, MRL from ZF/ROOT.ML.  These support forward proof, resolving a
 
lcp 
parents: 
0 
diff
changeset
 | 
602  | 
(*As above, but for rule lists*)  | 
| 
252
 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 
wenzelm 
parents: 
229 
diff
changeset
 | 
603  | 
fun rlss MRL bottom_rls =  | 
| 
11
 
d0e17c42dbb4
Added MRS, MRL from ZF/ROOT.ML.  These support forward proof, resolving a
 
lcp 
parents: 
0 
diff
changeset
 | 
604  | 
let fun rs_aux i [] = bottom_rls  | 
| 
252
 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 
wenzelm 
parents: 
229 
diff
changeset
 | 
605  | 
| rs_aux i (rls::rlss) = rls RLN (i, rs_aux (i+1) rlss)  | 
| 
11
 
d0e17c42dbb4
Added MRS, MRL from ZF/ROOT.ML.  These support forward proof, resolving a
 
lcp 
parents: 
0 
diff
changeset
 | 
606  | 
in rs_aux 1 rlss end;  | 
| 
 
d0e17c42dbb4
Added MRS, MRL from ZF/ROOT.ML.  These support forward proof, resolving a
 
lcp 
parents: 
0 
diff
changeset
 | 
607  | 
|
| 
9288
 
06a55195741b
infix 'OF' is a version of 'MRS' with more appropriate argument order;
 
wenzelm 
parents: 
8605 
diff
changeset
 | 
608  | 
(*A version of MRS with more appropriate argument order*)  | 
| 
 
06a55195741b
infix 'OF' is a version of 'MRS' with more appropriate argument order;
 
wenzelm 
parents: 
8605 
diff
changeset
 | 
609  | 
fun bottom_rl OF rls = rls MRS bottom_rl;  | 
| 
 
06a55195741b
infix 'OF' is a version of 'MRS' with more appropriate argument order;
 
wenzelm 
parents: 
8605 
diff
changeset
 | 
610  | 
|
| 
252
 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 
wenzelm 
parents: 
229 
diff
changeset
 | 
611  | 
(*compose Q and [...,Qi,Q(i+1),...]==>R to [...,Q(i+1),...]==>R  | 
| 0 | 612  | 
with no lifting or renaming! Q may contain ==> or meta-quants  | 
613  | 
ALWAYS deletes premise i *)  | 
|
| 
252
 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 
wenzelm 
parents: 
229 
diff
changeset
 | 
614  | 
fun compose(tha,i,thb) =  | 
| 4270 | 615  | 
Seq.list_of (bicompose false (false,tha,0) i thb);  | 
| 0 | 616  | 
|
| 6946 | 617  | 
fun compose_single (tha,i,thb) =  | 
618  | 
(case compose (tha,i,thb) of  | 
|
619  | 
[th] => th  | 
|
620  | 
  | _ => raise THM ("compose: unique result expected", i, [tha,thb]));
 | 
|
621  | 
||
| 0 | 622  | 
(*compose Q and [Q1,Q2,...,Qk]==>R to [Q2,...,Qk]==>R getting unique result*)  | 
623  | 
fun tha COMP thb =  | 
|
624  | 
case compose(tha,1,thb) of  | 
|
| 
252
 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 
wenzelm 
parents: 
229 
diff
changeset
 | 
625  | 
[th] => th  | 
| 0 | 626  | 
      | _ =>   raise THM("COMP", 1, [tha,thb]);
 | 
627  | 
||
| 
13105
 
3d1e7a199bdc
use eq_thm_prop instead of slightly inadequate eq_thm;
 
wenzelm 
parents: 
12908 
diff
changeset
 | 
628  | 
|
| 4016 | 629  | 
(** theorem equality **)  | 
| 0 | 630  | 
|
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
15949 
diff
changeset
 | 
631  | 
(*True if the two theorems have the same theory.*)  | 
| 
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
15949 
diff
changeset
 | 
632  | 
val eq_thm_thy = eq_thy o pairself Thm.theory_of_thm;  | 
| 
13650
 
31bd2a8cdbe2
fixing the cut_tac method to work when there are no instantiations and the
 
paulson 
parents: 
13606 
diff
changeset
 | 
633  | 
|
| 
 
31bd2a8cdbe2
fixing the cut_tac method to work when there are no instantiations and the
 
paulson 
parents: 
13606 
diff
changeset
 | 
634  | 
(*True if the two theorems have the same prop field, ignoring hyps, der, etc.*)  | 
| 16720 | 635  | 
val eq_thm_prop = op aconv o pairself Thm.full_prop_of;  | 
| 0 | 636  | 
|
637  | 
(*Useful "distance" function for BEST_FIRST*)  | 
|
| 16720 | 638  | 
val size_of_thm = size_of_term o Thm.full_prop_of;  | 
| 0 | 639  | 
|
| 9829 | 640  | 
(*maintain lists of theorems --- preserving canonical order*)  | 
| 
13105
 
3d1e7a199bdc
use eq_thm_prop instead of slightly inadequate eq_thm;
 
wenzelm 
parents: 
12908 
diff
changeset
 | 
641  | 
fun del_rules rs rules = Library.gen_rems eq_thm_prop (rules, rs);  | 
| 9862 | 642  | 
fun add_rules rs rules = rs @ del_rules rs rules;  | 
| 12373 | 643  | 
val del_rule = del_rules o single;  | 
644  | 
val add_rule = add_rules o single;  | 
|
| 
13105
 
3d1e7a199bdc
use eq_thm_prop instead of slightly inadequate eq_thm;
 
wenzelm 
parents: 
12908 
diff
changeset
 | 
645  | 
fun merge_rules (rules1, rules2) = gen_merge_lists' eq_thm_prop rules1 rules2;  | 
| 9829 | 646  | 
|
| 
1194
 
563ecd14c1d8
Added weak_eq_thm and forall_intr_vars (thanks to Mark Staples)
 
lcp 
parents: 
952 
diff
changeset
 | 
647  | 
(** Mark Staples's weaker version of eq_thm: ignores variable renaming and  | 
| 
 
563ecd14c1d8
Added weak_eq_thm and forall_intr_vars (thanks to Mark Staples)
 
lcp 
parents: 
952 
diff
changeset
 | 
648  | 
(some) type variable renaming **)  | 
| 
 
563ecd14c1d8
Added weak_eq_thm and forall_intr_vars (thanks to Mark Staples)
 
lcp 
parents: 
952 
diff
changeset
 | 
649  | 
|
| 
 
563ecd14c1d8
Added weak_eq_thm and forall_intr_vars (thanks to Mark Staples)
 
lcp 
parents: 
952 
diff
changeset
 | 
650  | 
(* Can't use term_vars, because it sorts the resulting list of variable names.  | 
| 
 
563ecd14c1d8
Added weak_eq_thm and forall_intr_vars (thanks to Mark Staples)
 
lcp 
parents: 
952 
diff
changeset
 | 
651  | 
We instead need the unique list noramlised by the order of appearance  | 
| 
 
563ecd14c1d8
Added weak_eq_thm and forall_intr_vars (thanks to Mark Staples)
 
lcp 
parents: 
952 
diff
changeset
 | 
652  | 
in the term. *)  | 
| 
 
563ecd14c1d8
Added weak_eq_thm and forall_intr_vars (thanks to Mark Staples)
 
lcp 
parents: 
952 
diff
changeset
 | 
653  | 
fun term_vars' (t as Var(v,T)) = [t]  | 
| 
 
563ecd14c1d8
Added weak_eq_thm and forall_intr_vars (thanks to Mark Staples)
 
lcp 
parents: 
952 
diff
changeset
 | 
654  | 
| term_vars' (Abs(_,_,b)) = term_vars' b  | 
| 
 
563ecd14c1d8
Added weak_eq_thm and forall_intr_vars (thanks to Mark Staples)
 
lcp 
parents: 
952 
diff
changeset
 | 
655  | 
| term_vars' (f$a) = (term_vars' f) @ (term_vars' a)  | 
| 
 
563ecd14c1d8
Added weak_eq_thm and forall_intr_vars (thanks to Mark Staples)
 
lcp 
parents: 
952 
diff
changeset
 | 
656  | 
| term_vars' _ = [];  | 
| 
 
563ecd14c1d8
Added weak_eq_thm and forall_intr_vars (thanks to Mark Staples)
 
lcp 
parents: 
952 
diff
changeset
 | 
657  | 
|
| 
 
563ecd14c1d8
Added weak_eq_thm and forall_intr_vars (thanks to Mark Staples)
 
lcp 
parents: 
952 
diff
changeset
 | 
658  | 
fun forall_intr_vars th =  | 
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
15949 
diff
changeset
 | 
659  | 
  let val {prop,thy,...} = rep_thm th;
 | 
| 
1194
 
563ecd14c1d8
Added weak_eq_thm and forall_intr_vars (thanks to Mark Staples)
 
lcp 
parents: 
952 
diff
changeset
 | 
660  | 
val vars = distinct (term_vars' prop);  | 
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
15949 
diff
changeset
 | 
661  | 
in forall_intr_list (map (cterm_of thy) vars) th end;  | 
| 
1194
 
563ecd14c1d8
Added weak_eq_thm and forall_intr_vars (thanks to Mark Staples)
 
lcp 
parents: 
952 
diff
changeset
 | 
662  | 
|
| 
13105
 
3d1e7a199bdc
use eq_thm_prop instead of slightly inadequate eq_thm;
 
wenzelm 
parents: 
12908 
diff
changeset
 | 
663  | 
val weak_eq_thm = Thm.eq_thm o pairself (forall_intr_vars o freezeT);  | 
| 
1194
 
563ecd14c1d8
Added weak_eq_thm and forall_intr_vars (thanks to Mark Staples)
 
lcp 
parents: 
952 
diff
changeset
 | 
664  | 
|
| 
 
563ecd14c1d8
Added weak_eq_thm and forall_intr_vars (thanks to Mark Staples)
 
lcp 
parents: 
952 
diff
changeset
 | 
665  | 
|
| 0 | 666  | 
(*** Meta-Rewriting Rules ***)  | 
667  | 
||
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
15949 
diff
changeset
 | 
668  | 
fun read_prop s = read_cterm ProtoPure.thy (s, propT);  | 
| 
4610
 
b1322be47244
Tidying; rotate_prems; moved freeze_thaw from tactic.ML
 
paulson 
parents: 
4588 
diff
changeset
 | 
669  | 
|
| 9455 | 670  | 
fun store_thm name thm = hd (PureThy.smart_store_thms (name, [thm]));  | 
671  | 
fun store_standard_thm name thm = store_thm name (standard thm);  | 
|
| 12135 | 672  | 
fun store_thm_open name thm = hd (PureThy.smart_store_thms_open (name, [thm]));  | 
673  | 
fun store_standard_thm_open name thm = store_thm_open name (standard' thm);  | 
|
| 4016 | 674  | 
|
| 0 | 675  | 
val reflexive_thm =  | 
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
15949 
diff
changeset
 | 
676  | 
  let val cx = cterm_of ProtoPure.thy (Var(("x",0),TVar(("'a",0),[])))
 | 
| 12135 | 677  | 
in store_standard_thm_open "reflexive" (Thm.reflexive cx) end;  | 
| 0 | 678  | 
|
679  | 
val symmetric_thm =  | 
|
| 14854 | 680  | 
let val xy = read_prop "x == y"  | 
| 16595 | 681  | 
in store_standard_thm_open "symmetric" (Thm.implies_intr xy (Thm.symmetric (Thm.assume xy))) end;  | 
| 0 | 682  | 
|
683  | 
val transitive_thm =  | 
|
| 14854 | 684  | 
let val xy = read_prop "x == y"  | 
685  | 
val yz = read_prop "y == z"  | 
|
| 0 | 686  | 
val xythm = Thm.assume xy and yzthm = Thm.assume yz  | 
| 12135 | 687  | 
in store_standard_thm_open "transitive" (Thm.implies_intr yz (Thm.transitive xythm yzthm)) end;  | 
| 0 | 688  | 
|
| 4679 | 689  | 
fun symmetric_fun thm = thm RS symmetric_thm;  | 
690  | 
||
| 
11512
 
da3a96ab5630
Some basic rules are now stored with "open" derivations, to facilitate
 
berghofe 
parents: 
11163 
diff
changeset
 | 
691  | 
fun extensional eq =  | 
| 
 
da3a96ab5630
Some basic rules are now stored with "open" derivations, to facilitate
 
berghofe 
parents: 
11163 
diff
changeset
 | 
692  | 
let val eq' =  | 
| 
 
da3a96ab5630
Some basic rules are now stored with "open" derivations, to facilitate
 
berghofe 
parents: 
11163 
diff
changeset
 | 
693  | 
abstract_rule "x" (snd (Thm.dest_comb (fst (dest_equals (cprop_of eq))))) eq  | 
| 
 
da3a96ab5630
Some basic rules are now stored with "open" derivations, to facilitate
 
berghofe 
parents: 
11163 
diff
changeset
 | 
694  | 
in equal_elim (eta_conversion (cprop_of eq')) eq' end;  | 
| 
 
da3a96ab5630
Some basic rules are now stored with "open" derivations, to facilitate
 
berghofe 
parents: 
11163 
diff
changeset
 | 
695  | 
|
| 10414 | 696  | 
val imp_cong =  | 
697  | 
let  | 
|
698  | 
val ABC = read_prop "PROP A ==> PROP B == PROP C"  | 
|
699  | 
val AB = read_prop "PROP A ==> PROP B"  | 
|
700  | 
val AC = read_prop "PROP A ==> PROP C"  | 
|
701  | 
val A = read_prop "PROP A"  | 
|
702  | 
in  | 
|
| 12135 | 703  | 
store_standard_thm_open "imp_cong" (implies_intr ABC (equal_intr  | 
| 10414 | 704  | 
(implies_intr AB (implies_intr A  | 
705  | 
(equal_elim (implies_elim (assume ABC) (assume A))  | 
|
706  | 
(implies_elim (assume AB) (assume A)))))  | 
|
707  | 
(implies_intr AC (implies_intr A  | 
|
708  | 
(equal_elim (symmetric (implies_elim (assume ABC) (assume A)))  | 
|
709  | 
(implies_elim (assume AC) (assume A)))))))  | 
|
710  | 
end;  | 
|
711  | 
||
712  | 
val swap_prems_eq =  | 
|
713  | 
let  | 
|
714  | 
val ABC = read_prop "PROP A ==> PROP B ==> PROP C"  | 
|
715  | 
val BAC = read_prop "PROP B ==> PROP A ==> PROP C"  | 
|
716  | 
val A = read_prop "PROP A"  | 
|
717  | 
val B = read_prop "PROP B"  | 
|
718  | 
in  | 
|
| 12135 | 719  | 
store_standard_thm_open "swap_prems_eq" (equal_intr  | 
| 10414 | 720  | 
(implies_intr ABC (implies_intr B (implies_intr A  | 
721  | 
(implies_elim (implies_elim (assume ABC) (assume A)) (assume B)))))  | 
|
722  | 
(implies_intr BAC (implies_intr A (implies_intr B  | 
|
723  | 
(implies_elim (implies_elim (assume BAC) (assume B)) (assume A))))))  | 
|
724  | 
end;  | 
|
| 
229
 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 
lcp 
parents: 
214 
diff
changeset
 | 
725  | 
|
| 18468 | 726  | 
val imp_cong_rule = combination o combination (reflexive implies);  | 
| 0 | 727  | 
|
| 
15001
 
fb2141a9f8c0
Moved conversion rules from MetaSimplifier to Drule. refl_implies removed
 
skalberg 
parents: 
14854 
diff
changeset
 | 
728  | 
local  | 
| 
 
fb2141a9f8c0
Moved conversion rules from MetaSimplifier to Drule. refl_implies removed
 
skalberg 
parents: 
14854 
diff
changeset
 | 
729  | 
val dest_eq = dest_equals o cprop_of  | 
| 
 
fb2141a9f8c0
Moved conversion rules from MetaSimplifier to Drule. refl_implies removed
 
skalberg 
parents: 
14854 
diff
changeset
 | 
730  | 
val rhs_of = snd o dest_eq  | 
| 
 
fb2141a9f8c0
Moved conversion rules from MetaSimplifier to Drule. refl_implies removed
 
skalberg 
parents: 
14854 
diff
changeset
 | 
731  | 
in  | 
| 
 
fb2141a9f8c0
Moved conversion rules from MetaSimplifier to Drule. refl_implies removed
 
skalberg 
parents: 
14854 
diff
changeset
 | 
732  | 
fun beta_eta_conversion t =  | 
| 
 
fb2141a9f8c0
Moved conversion rules from MetaSimplifier to Drule. refl_implies removed
 
skalberg 
parents: 
14854 
diff
changeset
 | 
733  | 
let val thm = beta_conversion true t  | 
| 
 
fb2141a9f8c0
Moved conversion rules from MetaSimplifier to Drule. refl_implies removed
 
skalberg 
parents: 
14854 
diff
changeset
 | 
734  | 
in transitive thm (eta_conversion (rhs_of thm)) end  | 
| 
 
fb2141a9f8c0
Moved conversion rules from MetaSimplifier to Drule. refl_implies removed
 
skalberg 
parents: 
14854 
diff
changeset
 | 
735  | 
end;  | 
| 
 
fb2141a9f8c0
Moved conversion rules from MetaSimplifier to Drule. refl_implies removed
 
skalberg 
parents: 
14854 
diff
changeset
 | 
736  | 
|
| 15925 | 737  | 
fun eta_long_conversion ct = transitive (beta_eta_conversion ct)  | 
738  | 
(symmetric (beta_eta_conversion (cterm_fun (Pattern.eta_long []) ct)));  | 
|
739  | 
||
| 18337 | 740  | 
val abs_def =  | 
741  | 
let  | 
|
742  | 
fun contract_lhs th =  | 
|
743  | 
Thm.transitive (Thm.symmetric (beta_eta_conversion (fst (dest_equals (cprop_of th))))) th;  | 
|
744  | 
fun abstract cx = Thm.abstract_rule  | 
|
745  | 
(case Thm.term_of cx of Var ((x, _), _) => x | Free (x, _) => x | _ => "x") cx;  | 
|
746  | 
in  | 
|
747  | 
contract_lhs  | 
|
748  | 
#> `(snd o strip_comb o fst o dest_equals o cprop_of)  | 
|
749  | 
#-> fold_rev abstract  | 
|
750  | 
#> contract_lhs  | 
|
751  | 
end;  | 
|
752  | 
||
| 18468 | 753  | 
(*rewrite B in !!x1 ... xn. B*)  | 
| 18251 | 754  | 
fun forall_conv 0 cv ct = cv ct  | 
755  | 
| forall_conv n cv ct =  | 
|
| 18468 | 756  | 
(case try Thm.dest_comb ct of  | 
757  | 
NONE => cv ct  | 
|
758  | 
| SOME (A, B) =>  | 
|
759  | 
(case (term_of A, term_of B) of  | 
|
760  | 
            (Const ("all", _), Abs (x, _, _)) =>
 | 
|
761  | 
let val (v, B') = Thm.dest_abs (SOME (gensym "all_")) B in  | 
|
762  | 
Thm.combination (Thm.reflexive A)  | 
|
763  | 
(Thm.abstract_rule x v (forall_conv (n - 1) cv B'))  | 
|
764  | 
end  | 
|
765  | 
| _ => cv ct));  | 
|
766  | 
||
767  | 
(*rewrite B in A1 ==> ... ==> An ==> B*)  | 
|
768  | 
fun concl_conv 0 cv ct = cv ct  | 
|
769  | 
| concl_conv n cv ct =  | 
|
770  | 
(case try dest_implies ct of  | 
|
771  | 
NONE => cv ct  | 
|
772  | 
| SOME (A, B) => imp_cong_rule (reflexive A) (concl_conv (n - 1) cv B));  | 
|
| 
15001
 
fb2141a9f8c0
Moved conversion rules from MetaSimplifier to Drule. refl_implies removed
 
skalberg 
parents: 
14854 
diff
changeset
 | 
773  | 
|
| 18468 | 774  | 
(*rewrite the A's in A1 ==> ... ==> An ==> B*)  | 
775  | 
fun prems_conv 0 _ = reflexive  | 
|
776  | 
| prems_conv n cv =  | 
|
777  | 
let  | 
|
778  | 
fun conv i ct =  | 
|
779  | 
if i = n + 1 then reflexive ct  | 
|
780  | 
else  | 
|
781  | 
(case try dest_implies ct of  | 
|
782  | 
NONE => reflexive ct  | 
|
783  | 
| SOME (A, B) => imp_cong_rule (cv i A) (conv (i + 1) B));  | 
|
784  | 
in conv 1 end;  | 
|
785  | 
||
786  | 
(*rewrite the A's in A1 && ... && An*)  | 
|
787  | 
fun conjunction_conv 0 _ = reflexive  | 
|
788  | 
| conjunction_conv n cv =  | 
|
789  | 
let  | 
|
790  | 
fun conv i ct =  | 
|
791  | 
if i <> n andalso can Logic.dest_conjunction (term_of ct) then  | 
|
792  | 
forall_conv 1  | 
|
793  | 
(prems_conv 1 (K (prems_conv 2 (fn 1 => cv i | 2 => conv (i + 1))))) ct  | 
|
794  | 
else cv i ct;  | 
|
795  | 
in conv 1 end;  | 
|
796  | 
||
797  | 
||
798  | 
fun goals_conv pred cv = prems_conv ~1 (fn i => if pred i then cv else reflexive);  | 
|
| 
15001
 
fb2141a9f8c0
Moved conversion rules from MetaSimplifier to Drule. refl_implies removed
 
skalberg 
parents: 
14854 
diff
changeset
 | 
799  | 
fun fconv_rule cv th = equal_elim (cv (cprop_of th)) th;  | 
| 
 
fb2141a9f8c0
Moved conversion rules from MetaSimplifier to Drule. refl_implies removed
 
skalberg 
parents: 
14854 
diff
changeset
 | 
800  | 
|
| 18468 | 801  | 
|
| 15669 | 802  | 
(*** Some useful meta-theorems ***)  | 
| 0 | 803  | 
|
804  | 
(*The rule V/V, obtains assumption solving for eresolve_tac*)  | 
|
| 12135 | 805  | 
val asm_rl = store_standard_thm_open "asm_rl" (Thm.trivial (read_prop "PROP ?psi"));  | 
| 7380 | 806  | 
val _ = store_thm "_" asm_rl;  | 
| 0 | 807  | 
|
808  | 
(*Meta-level cut rule: [| V==>W; V |] ==> W *)  | 
|
| 4016 | 809  | 
val cut_rl =  | 
| 12135 | 810  | 
store_standard_thm_open "cut_rl"  | 
| 9455 | 811  | 
(Thm.trivial (read_prop "PROP ?psi ==> PROP ?theta"));  | 
| 0 | 812  | 
|
| 
252
 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 
wenzelm 
parents: 
229 
diff
changeset
 | 
813  | 
(*Generalized elim rule for one conclusion; cut_rl with reversed premises:  | 
| 0 | 814  | 
[| PROP V; PROP V ==> PROP W |] ==> PROP W *)  | 
815  | 
val revcut_rl =  | 
|
| 
4610
 
b1322be47244
Tidying; rotate_prems; moved freeze_thaw from tactic.ML
 
paulson 
parents: 
4588 
diff
changeset
 | 
816  | 
let val V = read_prop "PROP V"  | 
| 
 
b1322be47244
Tidying; rotate_prems; moved freeze_thaw from tactic.ML
 
paulson 
parents: 
4588 
diff
changeset
 | 
817  | 
and VW = read_prop "PROP V ==> PROP W";  | 
| 4016 | 818  | 
in  | 
| 12135 | 819  | 
store_standard_thm_open "revcut_rl"  | 
| 4016 | 820  | 
(implies_intr V (implies_intr VW (implies_elim (assume VW) (assume V))))  | 
| 0 | 821  | 
end;  | 
822  | 
||
| 668 | 823  | 
(*for deleting an unwanted assumption*)  | 
824  | 
val thin_rl =  | 
|
| 
4610
 
b1322be47244
Tidying; rotate_prems; moved freeze_thaw from tactic.ML
 
paulson 
parents: 
4588 
diff
changeset
 | 
825  | 
let val V = read_prop "PROP V"  | 
| 
 
b1322be47244
Tidying; rotate_prems; moved freeze_thaw from tactic.ML
 
paulson 
parents: 
4588 
diff
changeset
 | 
826  | 
and W = read_prop "PROP W";  | 
| 12135 | 827  | 
in store_standard_thm_open "thin_rl" (implies_intr V (implies_intr W (assume W))) end;  | 
| 668 | 828  | 
|
| 0 | 829  | 
(* (!!x. PROP ?V) == PROP ?V Allows removal of redundant parameters*)  | 
830  | 
val triv_forall_equality =  | 
|
| 
4610
 
b1322be47244
Tidying; rotate_prems; moved freeze_thaw from tactic.ML
 
paulson 
parents: 
4588 
diff
changeset
 | 
831  | 
let val V = read_prop "PROP V"  | 
| 
 
b1322be47244
Tidying; rotate_prems; moved freeze_thaw from tactic.ML
 
paulson 
parents: 
4588 
diff
changeset
 | 
832  | 
and QV = read_prop "!!x::'a. PROP V"  | 
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
15949 
diff
changeset
 | 
833  | 
      and x  = read_cterm ProtoPure.thy ("x", TypeInfer.logicT);
 | 
| 4016 | 834  | 
in  | 
| 12135 | 835  | 
store_standard_thm_open "triv_forall_equality"  | 
| 
11512
 
da3a96ab5630
Some basic rules are now stored with "open" derivations, to facilitate
 
berghofe 
parents: 
11163 
diff
changeset
 | 
836  | 
(equal_intr (implies_intr QV (forall_elim x (assume QV)))  | 
| 
 
da3a96ab5630
Some basic rules are now stored with "open" derivations, to facilitate
 
berghofe 
parents: 
11163 
diff
changeset
 | 
837  | 
(implies_intr V (forall_intr x (assume V))))  | 
| 0 | 838  | 
end;  | 
839  | 
||
| 1756 | 840  | 
(* (PROP ?PhiA ==> PROP ?PhiB ==> PROP ?Psi) ==>  | 
841  | 
(PROP ?PhiB ==> PROP ?PhiA ==> PROP ?Psi)  | 
|
842  | 
`thm COMP swap_prems_rl' swaps the first two premises of `thm'  | 
|
843  | 
*)  | 
|
844  | 
val swap_prems_rl =  | 
|
| 
4610
 
b1322be47244
Tidying; rotate_prems; moved freeze_thaw from tactic.ML
 
paulson 
parents: 
4588 
diff
changeset
 | 
845  | 
let val cmajor = read_prop "PROP PhiA ==> PROP PhiB ==> PROP Psi";  | 
| 1756 | 846  | 
val major = assume cmajor;  | 
| 
4610
 
b1322be47244
Tidying; rotate_prems; moved freeze_thaw from tactic.ML
 
paulson 
parents: 
4588 
diff
changeset
 | 
847  | 
val cminor1 = read_prop "PROP PhiA";  | 
| 1756 | 848  | 
val minor1 = assume cminor1;  | 
| 
4610
 
b1322be47244
Tidying; rotate_prems; moved freeze_thaw from tactic.ML
 
paulson 
parents: 
4588 
diff
changeset
 | 
849  | 
val cminor2 = read_prop "PROP PhiB";  | 
| 1756 | 850  | 
val minor2 = assume cminor2;  | 
| 12135 | 851  | 
in store_standard_thm_open "swap_prems_rl"  | 
| 1756 | 852  | 
(implies_intr cmajor (implies_intr cminor2 (implies_intr cminor1  | 
853  | 
(implies_elim (implies_elim major minor1) minor2))))  | 
|
854  | 
end;  | 
|
855  | 
||
| 3653 | 856  | 
(* [| PROP ?phi ==> PROP ?psi; PROP ?psi ==> PROP ?phi |]  | 
857  | 
==> PROP ?phi == PROP ?psi  | 
|
| 8328 | 858  | 
Introduction rule for == as a meta-theorem.  | 
| 3653 | 859  | 
*)  | 
860  | 
val equal_intr_rule =  | 
|
| 
4610
 
b1322be47244
Tidying; rotate_prems; moved freeze_thaw from tactic.ML
 
paulson 
parents: 
4588 
diff
changeset
 | 
861  | 
let val PQ = read_prop "PROP phi ==> PROP psi"  | 
| 
 
b1322be47244
Tidying; rotate_prems; moved freeze_thaw from tactic.ML
 
paulson 
parents: 
4588 
diff
changeset
 | 
862  | 
and QP = read_prop "PROP psi ==> PROP phi"  | 
| 4016 | 863  | 
in  | 
| 12135 | 864  | 
store_standard_thm_open "equal_intr_rule"  | 
| 4016 | 865  | 
(implies_intr PQ (implies_intr QP (equal_intr (assume PQ) (assume QP))))  | 
| 3653 | 866  | 
end;  | 
867  | 
||
| 13368 | 868  | 
(* [| PROP ?phi == PROP ?psi; PROP ?phi |] ==> PROP ?psi *)  | 
869  | 
val equal_elim_rule1 =  | 
|
870  | 
let val eq = read_prop "PROP phi == PROP psi"  | 
|
871  | 
and P = read_prop "PROP phi"  | 
|
872  | 
in store_standard_thm_open "equal_elim_rule1"  | 
|
873  | 
(Thm.equal_elim (assume eq) (assume P) |> implies_intr_list [eq, P])  | 
|
874  | 
end;  | 
|
| 4285 | 875  | 
|
| 12297 | 876  | 
(* "[| PROP ?phi; PROP ?phi; PROP ?psi |] ==> PROP ?psi" *)  | 
877  | 
||
878  | 
val remdups_rl =  | 
|
879  | 
let val P = read_prop "PROP phi" and Q = read_prop "PROP psi";  | 
|
880  | 
in store_standard_thm_open "remdups_rl" (implies_intr_list [P, P, Q] (Thm.assume Q)) end;  | 
|
881  | 
||
882  | 
||
| 9554 | 883  | 
(*(PROP ?phi ==> (!!x. PROP ?psi(x))) == (!!x. PROP ?phi ==> PROP ?psi(x))  | 
| 12297 | 884  | 
Rewrite rule for HHF normalization.*)  | 
| 9554 | 885  | 
|
886  | 
val norm_hhf_eq =  | 
|
887  | 
let  | 
|
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
15949 
diff
changeset
 | 
888  | 
val cert = Thm.cterm_of ProtoPure.thy;  | 
| 14854 | 889  | 
    val aT = TFree ("'a", []);
 | 
| 9554 | 890  | 
val all = Term.all aT;  | 
891  | 
    val x = Free ("x", aT);
 | 
|
892  | 
    val phi = Free ("phi", propT);
 | 
|
893  | 
    val psi = Free ("psi", aT --> propT);
 | 
|
894  | 
||
895  | 
val cx = cert x;  | 
|
896  | 
val cphi = cert phi;  | 
|
897  | 
    val lhs = cert (Logic.mk_implies (phi, all $ Abs ("x", aT, psi $ Bound 0)));
 | 
|
898  | 
    val rhs = cert (all $ Abs ("x", aT, Logic.mk_implies (phi, psi $ Bound 0)));
 | 
|
899  | 
in  | 
|
900  | 
Thm.equal_intr  | 
|
901  | 
(Thm.implies_elim (Thm.assume lhs) (Thm.assume cphi)  | 
|
902  | 
|> Thm.forall_elim cx  | 
|
903  | 
|> Thm.implies_intr cphi  | 
|
904  | 
|> Thm.forall_intr cx  | 
|
905  | 
|> Thm.implies_intr lhs)  | 
|
906  | 
(Thm.implies_elim  | 
|
907  | 
(Thm.assume rhs |> Thm.forall_elim cx) (Thm.assume cphi)  | 
|
908  | 
|> Thm.forall_intr cx  | 
|
909  | 
|> Thm.implies_intr cphi  | 
|
910  | 
|> Thm.implies_intr rhs)  | 
|
| 12135 | 911  | 
|> store_standard_thm_open "norm_hhf_eq"  | 
| 9554 | 912  | 
end;  | 
913  | 
||
| 18179 | 914  | 
val norm_hhf_prop = Logic.dest_equals (Thm.prop_of norm_hhf_eq);  | 
915  | 
||
| 12800 | 916  | 
fun is_norm_hhf tm =  | 
917  | 
let  | 
|
918  | 
    fun is_norm (Const ("==>", _) $ _ $ (Const ("all", _) $ _)) = false
 | 
|
919  | 
| is_norm (t $ u) = is_norm t andalso is_norm u  | 
|
920  | 
| is_norm (Abs (_, _, t)) = is_norm t  | 
|
921  | 
| is_norm _ = true;  | 
|
922  | 
in is_norm (Pattern.beta_eta_contract tm) end;  | 
|
923  | 
||
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
15949 
diff
changeset
 | 
924  | 
fun norm_hhf thy t =  | 
| 12800 | 925  | 
if is_norm_hhf t then t  | 
| 18179 | 926  | 
else Pattern.rewrite_term thy [norm_hhf_prop] [] t;  | 
927  | 
||
| 12800 | 928  | 
|
| 9554 | 929  | 
|
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
15949 
diff
changeset
 | 
930  | 
(*** Instantiate theorem th, reading instantiations in theory thy ****)  | 
| 
8129
 
29e239c7b8c2
Thm.instantiate no longer normalizes, but Drule.instantiate does
 
paulson 
parents: 
8086 
diff
changeset
 | 
931  | 
|
| 
 
29e239c7b8c2
Thm.instantiate no longer normalizes, but Drule.instantiate does
 
paulson 
parents: 
8086 
diff
changeset
 | 
932  | 
(*Version that normalizes the result: Thm.instantiate no longer does that*)  | 
| 
 
29e239c7b8c2
Thm.instantiate no longer normalizes, but Drule.instantiate does
 
paulson 
parents: 
8086 
diff
changeset
 | 
933  | 
fun instantiate instpair th = Thm.instantiate instpair th COMP asm_rl;  | 
| 
 
29e239c7b8c2
Thm.instantiate no longer normalizes, but Drule.instantiate does
 
paulson 
parents: 
8086 
diff
changeset
 | 
934  | 
|
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
15949 
diff
changeset
 | 
935  | 
fun read_instantiate_sg' thy sinsts th =  | 
| 
8129
 
29e239c7b8c2
Thm.instantiate no longer normalizes, but Drule.instantiate does
 
paulson 
parents: 
8086 
diff
changeset
 | 
936  | 
let val ts = types_sorts th;  | 
| 15669 | 937  | 
val used = add_used th [];  | 
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
15949 
diff
changeset
 | 
938  | 
in instantiate (read_insts thy ts ts used sinsts) th end;  | 
| 15797 | 939  | 
|
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
15949 
diff
changeset
 | 
940  | 
fun read_instantiate_sg thy sinsts th =  | 
| 
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
15949 
diff
changeset
 | 
941  | 
read_instantiate_sg' thy (map (apfst Syntax.indexname) sinsts) th;  | 
| 
8129
 
29e239c7b8c2
Thm.instantiate no longer normalizes, but Drule.instantiate does
 
paulson 
parents: 
8086 
diff
changeset
 | 
942  | 
|
| 
 
29e239c7b8c2
Thm.instantiate no longer normalizes, but Drule.instantiate does
 
paulson 
parents: 
8086 
diff
changeset
 | 
943  | 
(*Instantiate theorem th, reading instantiations under theory of th*)  | 
| 
 
29e239c7b8c2
Thm.instantiate no longer normalizes, but Drule.instantiate does
 
paulson 
parents: 
8086 
diff
changeset
 | 
944  | 
fun read_instantiate sinsts th =  | 
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
15949 
diff
changeset
 | 
945  | 
read_instantiate_sg (Thm.theory_of_thm th) sinsts th;  | 
| 
8129
 
29e239c7b8c2
Thm.instantiate no longer normalizes, but Drule.instantiate does
 
paulson 
parents: 
8086 
diff
changeset
 | 
946  | 
|
| 15797 | 947  | 
fun read_instantiate' sinsts th =  | 
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
15949 
diff
changeset
 | 
948  | 
read_instantiate_sg' (Thm.theory_of_thm th) sinsts th;  | 
| 15797 | 949  | 
|
| 
8129
 
29e239c7b8c2
Thm.instantiate no longer normalizes, but Drule.instantiate does
 
paulson 
parents: 
8086 
diff
changeset
 | 
950  | 
|
| 
 
29e239c7b8c2
Thm.instantiate no longer normalizes, but Drule.instantiate does
 
paulson 
parents: 
8086 
diff
changeset
 | 
951  | 
(*Left-to-right replacements: tpairs = [...,(vi,ti),...].  | 
| 
 
29e239c7b8c2
Thm.instantiate no longer normalizes, but Drule.instantiate does
 
paulson 
parents: 
8086 
diff
changeset
 | 
952  | 
Instantiates distinct Vars by terms, inferring type instantiations. *)  | 
| 
 
29e239c7b8c2
Thm.instantiate no longer normalizes, but Drule.instantiate does
 
paulson 
parents: 
8086 
diff
changeset
 | 
953  | 
local  | 
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
15949 
diff
changeset
 | 
954  | 
fun add_types ((ct,cu), (thy,tye,maxidx)) =  | 
| 
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
15949 
diff
changeset
 | 
955  | 
    let val {thy=thyt, t=t, T= T, maxidx=maxt,...} = rep_cterm ct
 | 
| 
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
15949 
diff
changeset
 | 
956  | 
        and {thy=thyu, t=u, T= U, maxidx=maxu,...} = rep_cterm cu;
 | 
| 
8129
 
29e239c7b8c2
Thm.instantiate no longer normalizes, but Drule.instantiate does
 
paulson 
parents: 
8086 
diff
changeset
 | 
957  | 
val maxi = Int.max(maxidx, Int.max(maxt, maxu));  | 
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
15949 
diff
changeset
 | 
958  | 
val thy' = Theory.merge(thy, Theory.merge(thyt, thyu))  | 
| 
16949
 
ea65d75e0ce1
tuned gen_all, forall_elim_list, implies_intr_list, standard;
 
wenzelm 
parents: 
16861 
diff
changeset
 | 
959  | 
val (tye',maxi') = Sign.typ_unify thy' (T, U) (tye, maxi)  | 
| 10403 | 960  | 
          handle Type.TUNIFY => raise TYPE("Ill-typed instantiation", [T,U], [t,u])
 | 
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
15949 
diff
changeset
 | 
961  | 
in (thy', tye', maxi') end;  | 
| 
8129
 
29e239c7b8c2
Thm.instantiate no longer normalizes, but Drule.instantiate does
 
paulson 
parents: 
8086 
diff
changeset
 | 
962  | 
in  | 
| 
 
29e239c7b8c2
Thm.instantiate no longer normalizes, but Drule.instantiate does
 
paulson 
parents: 
8086 
diff
changeset
 | 
963  | 
fun cterm_instantiate ctpairs0 th =  | 
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
15949 
diff
changeset
 | 
964  | 
let val (thy,tye,_) = foldr add_types (Thm.theory_of_thm th, Vartab.empty, 0) ctpairs0  | 
| 18179 | 965  | 
fun instT(ct,cu) =  | 
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
15949 
diff
changeset
 | 
966  | 
let val inst = cterm_of thy o Envir.subst_TVars tye o term_of  | 
| 
14340
 
bc93ffa674cc
correction to cterm_instantiate by Christoph Leuth
 
paulson 
parents: 
14081 
diff
changeset
 | 
967  | 
in (inst ct, inst cu) end  | 
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
15949 
diff
changeset
 | 
968  | 
fun ctyp2 (ixn, (S, T)) = (ctyp_of thy (TVar (ixn, S)), ctyp_of thy T)  | 
| 
8406
 
a217b0cd304d
Type.unify and Type.typ_match now use Vartab instead of association lists.
 
berghofe 
parents: 
8365 
diff
changeset
 | 
969  | 
in instantiate (map ctyp2 (Vartab.dest tye), map instT ctpairs0) th end  | 
| 
8129
 
29e239c7b8c2
Thm.instantiate no longer normalizes, but Drule.instantiate does
 
paulson 
parents: 
8086 
diff
changeset
 | 
970  | 
handle TERM _ =>  | 
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
15949 
diff
changeset
 | 
971  | 
           raise THM("cterm_instantiate: incompatible theories",0,[th])
 | 
| 
8129
 
29e239c7b8c2
Thm.instantiate no longer normalizes, but Drule.instantiate does
 
paulson 
parents: 
8086 
diff
changeset
 | 
972  | 
| TYPE (msg, _, _) => raise THM(msg, 0, [th])  | 
| 
 
29e239c7b8c2
Thm.instantiate no longer normalizes, but Drule.instantiate does
 
paulson 
parents: 
8086 
diff
changeset
 | 
973  | 
end;  | 
| 
 
29e239c7b8c2
Thm.instantiate no longer normalizes, but Drule.instantiate does
 
paulson 
parents: 
8086 
diff
changeset
 | 
974  | 
|
| 
 
29e239c7b8c2
Thm.instantiate no longer normalizes, but Drule.instantiate does
 
paulson 
parents: 
8086 
diff
changeset
 | 
975  | 
|
| 
 
29e239c7b8c2
Thm.instantiate no longer normalizes, but Drule.instantiate does
 
paulson 
parents: 
8086 
diff
changeset
 | 
976  | 
(** Derived rules mainly for METAHYPS **)  | 
| 
 
29e239c7b8c2
Thm.instantiate no longer normalizes, but Drule.instantiate does
 
paulson 
parents: 
8086 
diff
changeset
 | 
977  | 
|
| 
 
29e239c7b8c2
Thm.instantiate no longer normalizes, but Drule.instantiate does
 
paulson 
parents: 
8086 
diff
changeset
 | 
978  | 
(*Given the term "a", takes (%x.t)==(%x.u) to t[a/x]==u[a/x]*)  | 
| 
 
29e239c7b8c2
Thm.instantiate no longer normalizes, but Drule.instantiate does
 
paulson 
parents: 
8086 
diff
changeset
 | 
979  | 
fun equal_abs_elim ca eqth =  | 
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
15949 
diff
changeset
 | 
980  | 
  let val {thy=thya, t=a, ...} = rep_cterm ca
 | 
| 
8129
 
29e239c7b8c2
Thm.instantiate no longer normalizes, but Drule.instantiate does
 
paulson 
parents: 
8086 
diff
changeset
 | 
981  | 
and combth = combination eqth (reflexive ca)  | 
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
15949 
diff
changeset
 | 
982  | 
      val {thy,prop,...} = rep_thm eqth
 | 
| 
8129
 
29e239c7b8c2
Thm.instantiate no longer normalizes, but Drule.instantiate does
 
paulson 
parents: 
8086 
diff
changeset
 | 
983  | 
val (abst,absu) = Logic.dest_equals prop  | 
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
15949 
diff
changeset
 | 
984  | 
val cterm = cterm_of (Theory.merge (thy,thya))  | 
| 10414 | 985  | 
in transitive (symmetric (beta_conversion false (cterm (abst$a))))  | 
986  | 
(transitive combth (beta_conversion false (cterm (absu$a))))  | 
|
| 
8129
 
29e239c7b8c2
Thm.instantiate no longer normalizes, but Drule.instantiate does
 
paulson 
parents: 
8086 
diff
changeset
 | 
987  | 
end  | 
| 
 
29e239c7b8c2
Thm.instantiate no longer normalizes, but Drule.instantiate does
 
paulson 
parents: 
8086 
diff
changeset
 | 
988  | 
  handle THM _ => raise THM("equal_abs_elim", 0, [eqth]);
 | 
| 
 
29e239c7b8c2
Thm.instantiate no longer normalizes, but Drule.instantiate does
 
paulson 
parents: 
8086 
diff
changeset
 | 
989  | 
|
| 
 
29e239c7b8c2
Thm.instantiate no longer normalizes, but Drule.instantiate does
 
paulson 
parents: 
8086 
diff
changeset
 | 
990  | 
(*Calling equal_abs_elim with multiple terms*)  | 
| 
15574
 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 
skalberg 
parents: 
15570 
diff
changeset
 | 
991  | 
fun equal_abs_elim_list cts th = foldr (uncurry equal_abs_elim) th (rev cts);  | 
| 
8129
 
29e239c7b8c2
Thm.instantiate no longer normalizes, but Drule.instantiate does
 
paulson 
parents: 
8086 
diff
changeset
 | 
992  | 
|
| 
 
29e239c7b8c2
Thm.instantiate no longer normalizes, but Drule.instantiate does
 
paulson 
parents: 
8086 
diff
changeset
 | 
993  | 
|
| 18025 | 994  | 
(** protected propositions **)  | 
| 4789 | 995  | 
|
996  | 
local  | 
|
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
15949 
diff
changeset
 | 
997  | 
val cert = Thm.cterm_of ProtoPure.thy;  | 
| 18025 | 998  | 
  val A = cert (Free ("A", propT));
 | 
999  | 
val prop_def = #1 (freeze_thaw ProtoPure.prop_def);  | 
|
| 4789 | 1000  | 
in  | 
| 18025 | 1001  | 
val protect = Thm.capply (cert Logic.protectC);  | 
1002  | 
val protectI = store_thm "protectI" (kind_rule internalK (standard  | 
|
1003  | 
(Thm.equal_elim (Thm.symmetric prop_def) (Thm.assume A))));  | 
|
1004  | 
val protectD = store_thm "protectD" (kind_rule internalK (standard  | 
|
1005  | 
(Thm.equal_elim prop_def (Thm.assume (protect A)))));  | 
|
| 18179 | 1006  | 
val protect_cong = store_standard_thm_open "protect_cong" (Thm.reflexive (protect A));  | 
| 4789 | 1007  | 
end;  | 
1008  | 
||
| 18025 | 1009  | 
fun implies_intr_protected asms th =  | 
| 18118 | 1010  | 
let val asms' = map protect asms in  | 
1011  | 
implies_elim_list  | 
|
1012  | 
(implies_intr_list asms th)  | 
|
1013  | 
(map (fn asm' => Thm.assume asm' RS protectD) asms')  | 
|
1014  | 
|> implies_intr_list asms'  | 
|
1015  | 
end;  | 
|
| 11815 | 1016  | 
|
| 4789 | 1017  | 
|
| 5688 | 1018  | 
(** variations on instantiate **)  | 
| 4285 | 1019  | 
|
| 
8550
 
b44c185f8018
new meta-rule "inst", a shorthand for read_instantiate_sg
 
paulson 
parents: 
8496 
diff
changeset
 | 
1020  | 
(*shorthand for instantiating just one variable in the current theory*)  | 
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
15949 
diff
changeset
 | 
1021  | 
fun inst x t = read_instantiate_sg (the_context()) [(x,t)];  | 
| 
8550
 
b44c185f8018
new meta-rule "inst", a shorthand for read_instantiate_sg
 
paulson 
parents: 
8496 
diff
changeset
 | 
1022  | 
|
| 
 
b44c185f8018
new meta-rule "inst", a shorthand for read_instantiate_sg
 
paulson 
parents: 
8496 
diff
changeset
 | 
1023  | 
|
| 4285 | 1024  | 
(* instantiate by left-to-right occurrence of variables *)  | 
1025  | 
||
1026  | 
fun instantiate' cTs cts thm =  | 
|
1027  | 
let  | 
|
1028  | 
fun err msg =  | 
|
1029  | 
      raise TYPE ("instantiate': " ^ msg,
 | 
|
| 15570 | 1030  | 
List.mapPartial (Option.map Thm.typ_of) cTs,  | 
1031  | 
List.mapPartial (Option.map Thm.term_of) cts);  | 
|
| 4285 | 1032  | 
|
1033  | 
fun inst_of (v, ct) =  | 
|
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
15949 
diff
changeset
 | 
1034  | 
(Thm.cterm_of (Thm.theory_of_cterm ct) (Var v), ct)  | 
| 4285 | 1035  | 
handle TYPE (msg, _, _) => err msg;  | 
1036  | 
||
| 15797 | 1037  | 
fun tyinst_of (v, cT) =  | 
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
15949 
diff
changeset
 | 
1038  | 
(Thm.ctyp_of (Thm.theory_of_ctyp cT) (TVar v), cT)  | 
| 15797 | 1039  | 
handle TYPE (msg, _, _) => err msg;  | 
1040  | 
||
| 4285 | 1041  | 
fun zip_vars _ [] = []  | 
| 15531 | 1042  | 
| zip_vars (_ :: vs) (NONE :: opt_ts) = zip_vars vs opt_ts  | 
1043  | 
| zip_vars (v :: vs) (SOME t :: opt_ts) = (v, t) :: zip_vars vs opt_ts  | 
|
| 4285 | 1044  | 
| zip_vars [] _ = err "more instantiations than variables in thm";  | 
1045  | 
||
1046  | 
(*instantiate types first!*)  | 
|
1047  | 
val thm' =  | 
|
1048  | 
if forall is_none cTs then thm  | 
|
| 15797 | 1049  | 
else Thm.instantiate (map tyinst_of (zip_vars (tvars_of thm) cTs), []) thm;  | 
| 4285 | 1050  | 
in  | 
1051  | 
if forall is_none cts then thm'  | 
|
1052  | 
else Thm.instantiate ([], map inst_of (zip_vars (vars_of thm') cts)) thm'  | 
|
1053  | 
end;  | 
|
1054  | 
||
1055  | 
||
| 14081 | 1056  | 
|
1057  | 
(** renaming of bound variables **)  | 
|
1058  | 
||
1059  | 
(* replace bound variables x_i in thm by y_i *)  | 
|
1060  | 
(* where vs = [(x_1, y_1), ..., (x_n, y_n)] *)  | 
|
1061  | 
||
1062  | 
fun rename_bvars [] thm = thm  | 
|
1063  | 
| rename_bvars vs thm =  | 
|
1064  | 
let  | 
|
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
15949 
diff
changeset
 | 
1065  | 
      val {thy, prop, ...} = rep_thm thm;
 | 
| 17325 | 1066  | 
fun ren (Abs (x, T, t)) = Abs (AList.lookup (op =) vs x |> the_default x, T, ren t)  | 
| 14081 | 1067  | 
| ren (t $ u) = ren t $ ren u  | 
1068  | 
| ren t = t;  | 
|
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
15949 
diff
changeset
 | 
1069  | 
in equal_elim (reflexive (cterm_of thy (ren prop))) thm end;  | 
| 14081 | 1070  | 
|
1071  | 
||
1072  | 
(* renaming in left-to-right order *)  | 
|
1073  | 
||
1074  | 
fun rename_bvars' xs thm =  | 
|
1075  | 
let  | 
|
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
15949 
diff
changeset
 | 
1076  | 
    val {thy, prop, ...} = rep_thm thm;
 | 
| 14081 | 1077  | 
fun rename [] t = ([], t)  | 
1078  | 
| rename (x' :: xs) (Abs (x, T, t)) =  | 
|
1079  | 
let val (xs', t') = rename xs t  | 
|
| 15570 | 1080  | 
in (xs', Abs (getOpt (x',x), T, t')) end  | 
| 14081 | 1081  | 
| rename xs (t $ u) =  | 
1082  | 
let  | 
|
1083  | 
val (xs', t') = rename xs t;  | 
|
1084  | 
val (xs'', u') = rename xs' u  | 
|
1085  | 
in (xs'', t' $ u') end  | 
|
1086  | 
| rename xs t = (xs, t);  | 
|
1087  | 
in case rename xs prop of  | 
|
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
15949 
diff
changeset
 | 
1088  | 
([], prop') => equal_elim (reflexive (cterm_of thy prop')) thm  | 
| 14081 | 1089  | 
| _ => error "More names than abstractions in theorem"  | 
1090  | 
end;  | 
|
1091  | 
||
1092  | 
||
1093  | 
||
| 5688 | 1094  | 
(* unvarify(T) *)  | 
1095  | 
||
1096  | 
(*assume thm in standard form, i.e. no frees, 0 var indexes*)  | 
|
1097  | 
||
1098  | 
fun unvarifyT thm =  | 
|
1099  | 
let  | 
|
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
15949 
diff
changeset
 | 
1100  | 
val cT = Thm.ctyp_of (Thm.theory_of_thm thm);  | 
| 15531 | 1101  | 
val tfrees = map (fn ((x, _), S) => SOME (cT (TFree (x, S)))) (tvars_of thm);  | 
| 5688 | 1102  | 
in instantiate' tfrees [] thm end;  | 
1103  | 
||
1104  | 
fun unvarify raw_thm =  | 
|
1105  | 
let  | 
|
1106  | 
val thm = unvarifyT raw_thm;  | 
|
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
15949 
diff
changeset
 | 
1107  | 
val ct = Thm.cterm_of (Thm.theory_of_thm thm);  | 
| 15531 | 1108  | 
val frees = map (fn ((x, _), T) => SOME (ct (Free (x, T)))) (vars_of thm);  | 
| 5688 | 1109  | 
in instantiate' [] frees thm end;  | 
1110  | 
||
1111  | 
||
| 8605 | 1112  | 
(* tvars_intr_list *)  | 
1113  | 
||
1114  | 
fun tvars_intr_list tfrees thm =  | 
|
| 18129 | 1115  | 
apfst (map (fn ((s, S), ixn) => (s, (ixn, S)))) (Thm.varifyT'  | 
| 15797 | 1116  | 
(gen_rems (op = o apfst fst) (tfrees_of thm, tfrees)) thm);  | 
| 8605 | 1117  | 
|
1118  | 
||
| 6435 | 1119  | 
(* increment var indexes *)  | 
1120  | 
||
| 18025 | 1121  | 
fun incr_indexes th = Thm.incr_indexes (#maxidx (Thm.rep_thm th) + 1);  | 
1122  | 
||
| 6435 | 1123  | 
fun incr_indexes_wrt is cTs cts thms =  | 
1124  | 
let  | 
|
1125  | 
val maxidx =  | 
|
| 15570 | 1126  | 
Library.foldl Int.max (~1, is @  | 
| 6435 | 1127  | 
map (maxidx_of_typ o #T o Thm.rep_ctyp) cTs @  | 
1128  | 
map (#maxidx o Thm.rep_cterm) cts @  | 
|
1129  | 
map (#maxidx o Thm.rep_thm) thms);  | 
|
| 10414 | 1130  | 
in Thm.incr_indexes (maxidx + 1) end;  | 
| 6435 | 1131  | 
|
1132  | 
||
| 8328 | 1133  | 
(* freeze_all *)  | 
1134  | 
||
1135  | 
(*freeze all (T)Vars; assumes thm in standard form*)  | 
|
1136  | 
||
1137  | 
fun freeze_all_TVars thm =  | 
|
1138  | 
(case tvars_of thm of  | 
|
1139  | 
[] => thm  | 
|
1140  | 
| tvars =>  | 
|
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
15949 
diff
changeset
 | 
1141  | 
let val cert = Thm.ctyp_of (Thm.theory_of_thm thm)  | 
| 15531 | 1142  | 
in instantiate' (map (fn ((x, _), S) => SOME (cert (TFree (x, S)))) tvars) [] thm end);  | 
| 8328 | 1143  | 
|
1144  | 
fun freeze_all_Vars thm =  | 
|
1145  | 
(case vars_of thm of  | 
|
1146  | 
[] => thm  | 
|
1147  | 
| vars =>  | 
|
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
15949 
diff
changeset
 | 
1148  | 
let val cert = Thm.cterm_of (Thm.theory_of_thm thm)  | 
| 15531 | 1149  | 
in instantiate' [] (map (fn ((x, _), T) => SOME (cert (Free (x, T)))) vars) thm end);  | 
| 8328 | 1150  | 
|
1151  | 
val freeze_all = freeze_all_Vars o freeze_all_TVars;  | 
|
1152  | 
||
1153  | 
||
| 11975 | 1154  | 
|
| 18225 | 1155  | 
(** multi_resolve **)  | 
1156  | 
||
1157  | 
local  | 
|
1158  | 
||
1159  | 
fun res th i rule =  | 
|
1160  | 
Thm.biresolution false [(false, th)] i rule handle THM _ => Seq.empty;  | 
|
1161  | 
||
1162  | 
fun multi_res _ [] rule = Seq.single rule  | 
|
1163  | 
| multi_res i (th :: ths) rule = Seq.maps (res th i) (multi_res (i + 1) ths rule);  | 
|
1164  | 
||
1165  | 
in  | 
|
1166  | 
||
1167  | 
val multi_resolve = multi_res 1;  | 
|
1168  | 
fun multi_resolves facts rules = Seq.maps (multi_resolve facts) (Seq.of_list rules);  | 
|
1169  | 
||
1170  | 
end;  | 
|
1171  | 
||
1172  | 
||
1173  | 
||
| 11975 | 1174  | 
(** meta-level conjunction **)  | 
1175  | 
||
1176  | 
local  | 
|
1177  | 
val A = read_prop "PROP A";  | 
|
1178  | 
val B = read_prop "PROP B";  | 
|
1179  | 
val C = read_prop "PROP C";  | 
|
1180  | 
val ABC = read_prop "PROP A ==> PROP B ==> PROP C";  | 
|
1181  | 
||
1182  | 
val proj1 =  | 
|
1183  | 
forall_intr_list [A, B] (implies_intr_list [A, B] (Thm.assume A))  | 
|
1184  | 
|> forall_elim_vars 0;  | 
|
1185  | 
||
1186  | 
val proj2 =  | 
|
1187  | 
forall_intr_list [A, B] (implies_intr_list [A, B] (Thm.assume B))  | 
|
1188  | 
|> forall_elim_vars 0;  | 
|
1189  | 
||
1190  | 
val conj_intr_rule =  | 
|
1191  | 
forall_intr_list [A, B] (implies_intr_list [A, B]  | 
|
1192  | 
(Thm.forall_intr C (Thm.implies_intr ABC  | 
|
1193  | 
(implies_elim_list (Thm.assume ABC) [Thm.assume A, Thm.assume B]))))  | 
|
1194  | 
|> forall_elim_vars 0;  | 
|
1195  | 
in  | 
|
1196  | 
||
| 18025 | 1197  | 
fun conj_intr tha thb = thb COMP (tha COMP incr_indexes_wrt [] [] [] [tha, thb] conj_intr_rule);  | 
| 
12756
 
08bf3d911968
conj_intr_list and conj_elim_precise: cover 0 conjuncts;
 
wenzelm 
parents: 
12725 
diff
changeset
 | 
1198  | 
|
| 
 
08bf3d911968
conj_intr_list and conj_elim_precise: cover 0 conjuncts;
 
wenzelm 
parents: 
12725 
diff
changeset
 | 
1199  | 
fun conj_intr_list [] = asm_rl  | 
| 
 
08bf3d911968
conj_intr_list and conj_elim_precise: cover 0 conjuncts;
 
wenzelm 
parents: 
12725 
diff
changeset
 | 
1200  | 
| conj_intr_list ths = foldr1 (uncurry conj_intr) ths;  | 
| 11975 | 1201  | 
|
1202  | 
fun conj_elim th =  | 
|
1203  | 
let val th' = forall_elim_var (#maxidx (Thm.rep_thm th) + 1) th  | 
|
| 18025 | 1204  | 
in (incr_indexes th' proj1 COMP th', incr_indexes th' proj2 COMP th') end;  | 
| 11975 | 1205  | 
|
| 
18498
 
466351242c6f
conj_elim_precise: proper treatment of nested conjunctions;
 
wenzelm 
parents: 
18468 
diff
changeset
 | 
1206  | 
(*((A && B) && C) && D && E -- flat*)  | 
| 11975 | 1207  | 
fun conj_elim_list th =  | 
1208  | 
let val (th1, th2) = conj_elim th  | 
|
1209  | 
in conj_elim_list th1 @ conj_elim_list th2 end handle THM _ => [th];  | 
|
1210  | 
||
| 
18498
 
466351242c6f
conj_elim_precise: proper treatment of nested conjunctions;
 
wenzelm 
parents: 
18468 
diff
changeset
 | 
1211  | 
(*(A1 && B1 && C1) && (A2 && B2 && C2 && D2) && A3 && B3 -- improper*)  | 
| 
 
466351242c6f
conj_elim_precise: proper treatment of nested conjunctions;
 
wenzelm 
parents: 
18468 
diff
changeset
 | 
1212  | 
fun conj_elim_precise spans =  | 
| 
 
466351242c6f
conj_elim_precise: proper treatment of nested conjunctions;
 
wenzelm 
parents: 
18468 
diff
changeset
 | 
1213  | 
let  | 
| 
 
466351242c6f
conj_elim_precise: proper treatment of nested conjunctions;
 
wenzelm 
parents: 
18468 
diff
changeset
 | 
1214  | 
fun elim 0 _ = []  | 
| 
 
466351242c6f
conj_elim_precise: proper treatment of nested conjunctions;
 
wenzelm 
parents: 
18468 
diff
changeset
 | 
1215  | 
| elim 1 th = [th]  | 
| 
 
466351242c6f
conj_elim_precise: proper treatment of nested conjunctions;
 
wenzelm 
parents: 
18468 
diff
changeset
 | 
1216  | 
| elim n th =  | 
| 
 
466351242c6f
conj_elim_precise: proper treatment of nested conjunctions;
 
wenzelm 
parents: 
18468 
diff
changeset
 | 
1217  | 
let val (th1, th2) = conj_elim th  | 
| 
 
466351242c6f
conj_elim_precise: proper treatment of nested conjunctions;
 
wenzelm 
parents: 
18468 
diff
changeset
 | 
1218  | 
in th1 :: elim (n - 1) th2 end;  | 
| 
 
466351242c6f
conj_elim_precise: proper treatment of nested conjunctions;
 
wenzelm 
parents: 
18468 
diff
changeset
 | 
1219  | 
fun elims (0 :: ns) ths = [] :: elims ns ths  | 
| 
 
466351242c6f
conj_elim_precise: proper treatment of nested conjunctions;
 
wenzelm 
parents: 
18468 
diff
changeset
 | 
1220  | 
| elims (n :: ns) (th :: ths) = elim n th :: elims ns ths  | 
| 
 
466351242c6f
conj_elim_precise: proper treatment of nested conjunctions;
 
wenzelm 
parents: 
18468 
diff
changeset
 | 
1221  | 
| elims _ _ = [];  | 
| 
 
466351242c6f
conj_elim_precise: proper treatment of nested conjunctions;
 
wenzelm 
parents: 
18468 
diff
changeset
 | 
1222  | 
in elims spans o elim (length (filter_out (equal 0) spans)) end;  | 
| 12135 | 1223  | 
|
1224  | 
val conj_intr_thm = store_standard_thm_open "conjunctionI"  | 
|
1225  | 
(implies_intr_list [A, B] (conj_intr (Thm.assume A) (Thm.assume B)));  | 
|
1226  | 
||
| 18206 | 1227  | 
end;  | 
| 18179 | 1228  | 
|
| 18206 | 1229  | 
fun conj_curry th =  | 
1230  | 
let  | 
|
1231  | 
    val {thy, maxidx, ...} = Thm.rep_thm th;
 | 
|
1232  | 
val n = Thm.nprems_of th;  | 
|
1233  | 
in  | 
|
1234  | 
if n < 2 then th  | 
|
1235  | 
else  | 
|
1236  | 
let  | 
|
1237  | 
val cert = Thm.cterm_of thy;  | 
|
1238  | 
        val As = map (fn i => Free ("A" ^ string_of_int i, propT)) (1 upto n);
 | 
|
1239  | 
        val B = Free ("B", propT);
 | 
|
1240  | 
val C = cert (Logic.mk_conjunction_list As);  | 
|
1241  | 
val D = cert (Logic.list_implies (As, B));  | 
|
1242  | 
val rule =  | 
|
1243  | 
implies_elim_list (Thm.assume D) (conj_elim_list (Thm.assume C))  | 
|
1244  | 
|> implies_intr_list [D, C]  | 
|
1245  | 
|> forall_intr_frees  | 
|
1246  | 
|> forall_elim_vars (maxidx + 1)  | 
|
1247  | 
in Thm.adjust_maxidx_thm (th COMP rule) end  | 
|
1248  | 
end;  | 
|
| 
252
 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 
wenzelm 
parents: 
229 
diff
changeset
 | 
1249  | 
|
| 11975 | 1250  | 
end;  | 
| 5903 | 1251  | 
|
1252  | 
structure BasicDrule: BASIC_DRULE = Drule;  | 
|
1253  | 
open BasicDrule;  |