author | wenzelm |
Fri, 26 May 2006 22:20:03 +0200 | |
changeset 19729 | cb9e2f0c7658 |
parent 19425 | e0d7d9373faf |
child 19775 | 06cb6743adf6 |
permissions | -rw-r--r-- |
9460 | 1 |
(* Title: Pure/logic.ML |
0 | 2 |
ID: $Id$ |
9460 | 3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
0 | 4 |
Copyright Cambridge University 1992 |
5 |
||
9460 | 6 |
Abstract syntax operations of the Pure meta-logic. |
0 | 7 |
*) |
8 |
||
9460 | 9 |
signature LOGIC = |
4345 | 10 |
sig |
18181 | 11 |
val is_all: term -> bool |
18762 | 12 |
val dest_all: term -> typ * term |
18181 | 13 |
val mk_equals: term * term -> term |
14 |
val dest_equals: term -> term * term |
|
15 |
val is_equals: term -> bool |
|
16 |
val mk_implies: term * term -> term |
|
17 |
val dest_implies: term -> term * term |
|
18 |
val is_implies: term -> bool |
|
19 |
val list_implies: term list * term -> term |
|
20 |
val strip_imp_prems: term -> term list |
|
21 |
val strip_imp_concl: term -> term |
|
22 |
val strip_prems: int * term list * term -> term list * term |
|
23 |
val count_prems: term * int -> int |
|
24 |
val nth_prem: int * term -> term |
|
19125
59b26248547b
simplified Pure conjunction, based on actual const;
wenzelm
parents:
19103
diff
changeset
|
25 |
val conjunction: term |
18181 | 26 |
val mk_conjunction: term * term -> term |
12757 | 27 |
val mk_conjunction_list: term list -> term |
18499 | 28 |
val mk_conjunction_list2: term list list -> term |
18469 | 29 |
val dest_conjunction: term -> term * term |
19425 | 30 |
val dest_conjunction_list: term -> term list |
18469 | 31 |
val dest_conjunctions: term -> term list |
18181 | 32 |
val strip_horn: term -> term list * term |
18938 | 33 |
val dest_type: term -> typ |
34 |
val const_of_class: class -> string |
|
35 |
val class_of_const: string -> class |
|
36 |
val mk_inclass: typ * class -> term |
|
37 |
val dest_inclass: term -> typ * class |
|
19406
410b9d9bf9a1
added mk_classrel, dest_classrel, mk_arities, dest_arity (from axclass.ML);
wenzelm
parents:
19391
diff
changeset
|
38 |
val mk_classrel: class * class -> term |
410b9d9bf9a1
added mk_classrel, dest_classrel, mk_arities, dest_arity (from axclass.ML);
wenzelm
parents:
19391
diff
changeset
|
39 |
val dest_classrel: term -> class * class |
410b9d9bf9a1
added mk_classrel, dest_classrel, mk_arities, dest_arity (from axclass.ML);
wenzelm
parents:
19391
diff
changeset
|
40 |
val mk_arities: string * sort list * sort -> term list |
410b9d9bf9a1
added mk_classrel, dest_classrel, mk_arities, dest_arity (from axclass.ML);
wenzelm
parents:
19391
diff
changeset
|
41 |
val dest_arity: term -> string * sort list * class |
18938 | 42 |
val dest_def: Pretty.pp -> (term -> bool) -> (string -> bool) -> (string -> bool) -> |
43 |
term -> (term * term) * term |
|
44 |
val abs_def: term -> term * term |
|
18181 | 45 |
val mk_cond_defpair: term list -> term * term -> string * term |
46 |
val mk_defpair: term * term -> string * term |
|
47 |
val mk_type: typ -> term |
|
48 |
val protectC: term |
|
49 |
val protect: term -> term |
|
50 |
val unprotect: term -> term |
|
51 |
val occs: term * term -> bool |
|
52 |
val close_form: term -> term |
|
18938 | 53 |
val combound: term * int * int -> term |
54 |
val rlist_abs: (string * typ) list * term -> term |
|
18181 | 55 |
val incr_indexes: typ list * int -> term -> term |
56 |
val incr_tvar: int -> typ -> typ |
|
57 |
val lift_abs: int -> term -> term -> term |
|
58 |
val lift_all: int -> term -> term -> term |
|
59 |
val strip_assums_hyp: term -> term list |
|
9460 | 60 |
val strip_assums_concl: term -> term |
18181 | 61 |
val strip_params: term -> (string * typ) list |
62 |
val has_meta_prems: term -> int -> bool |
|
63 |
val flatten_params: int -> term -> term |
|
64 |
val auto_rename: bool ref |
|
65 |
val set_rename_prefix: string -> unit |
|
0 | 66 |
val list_rename_params: string list * term -> term |
18181 | 67 |
val assum_pairs: int * term -> (term*term)list |
68 |
val varify: term -> term |
|
69 |
val unvarify: term -> term |
|
70 |
val get_goal: term -> int -> term |
|
71 |
val goal_params: term -> int -> term * term list |
|
72 |
val prems_of_goal: term -> int -> term list |
|
73 |
val concl_of_goal: term -> int -> term |
|
4345 | 74 |
end; |
0 | 75 |
|
1500 | 76 |
structure Logic : LOGIC = |
0 | 77 |
struct |
398
41f279b477e2
added mk_type, dest_type, mk_inclass, dest_inclass (for axclasses);
wenzelm
parents:
210
diff
changeset
|
78 |
|
4345 | 79 |
|
0 | 80 |
(*** Abstract syntax operations on the meta-connectives ***) |
81 |
||
5041
a1d0a6d555cd
Goals may now contain assumptions, which are not returned.
nipkow
parents:
4822
diff
changeset
|
82 |
(** all **) |
a1d0a6d555cd
Goals may now contain assumptions, which are not returned.
nipkow
parents:
4822
diff
changeset
|
83 |
|
a1d0a6d555cd
Goals may now contain assumptions, which are not returned.
nipkow
parents:
4822
diff
changeset
|
84 |
fun is_all (Const ("all", _) $ _) = true |
a1d0a6d555cd
Goals may now contain assumptions, which are not returned.
nipkow
parents:
4822
diff
changeset
|
85 |
| is_all _ = false; |
a1d0a6d555cd
Goals may now contain assumptions, which are not returned.
nipkow
parents:
4822
diff
changeset
|
86 |
|
18762 | 87 |
fun dest_all (Const ("all", Type ("fun", [Type ("fun", [T, _]), _])) $ A) = (T, A) |
88 |
| dest_all t = raise TERM ("dest_all", [t]); |
|
89 |
||
90 |
||
5041
a1d0a6d555cd
Goals may now contain assumptions, which are not returned.
nipkow
parents:
4822
diff
changeset
|
91 |
|
0 | 92 |
(** equality **) |
93 |
||
1835 | 94 |
(*Make an equality. DOES NOT CHECK TYPE OF u*) |
64
0bbe5d86cb38
logic/mk_equals,mk_flexpair: now calls fastype_of instead of type_of.
lcp
parents:
0
diff
changeset
|
95 |
fun mk_equals(t,u) = equals(fastype_of t) $ t $ u; |
0 | 96 |
|
97 |
fun dest_equals (Const("==",_) $ t $ u) = (t,u) |
|
98 |
| dest_equals t = raise TERM("dest_equals", [t]); |
|
99 |
||
637 | 100 |
fun is_equals (Const ("==", _) $ _ $ _) = true |
101 |
| is_equals _ = false; |
|
102 |
||
103 |
||
0 | 104 |
(** implies **) |
105 |
||
106 |
fun mk_implies(A,B) = implies $ A $ B; |
|
107 |
||
108 |
fun dest_implies (Const("==>",_) $ A $ B) = (A,B) |
|
109 |
| dest_implies A = raise TERM("dest_implies", [A]); |
|
110 |
||
5041
a1d0a6d555cd
Goals may now contain assumptions, which are not returned.
nipkow
parents:
4822
diff
changeset
|
111 |
fun is_implies (Const ("==>", _) $ _ $ _) = true |
a1d0a6d555cd
Goals may now contain assumptions, which are not returned.
nipkow
parents:
4822
diff
changeset
|
112 |
| is_implies _ = false; |
a1d0a6d555cd
Goals may now contain assumptions, which are not returned.
nipkow
parents:
4822
diff
changeset
|
113 |
|
4822 | 114 |
|
0 | 115 |
(** nested implications **) |
116 |
||
117 |
(* [A1,...,An], B goes to A1==>...An==>B *) |
|
18181 | 118 |
fun list_implies ([], B) = B |
119 |
| list_implies (A::As, B) = implies $ A $ list_implies(As,B); |
|
0 | 120 |
|
121 |
(* A1==>...An==>B goes to [A1,...,An], where B is not an implication *) |
|
122 |
fun strip_imp_prems (Const("==>", _) $ A $ B) = A :: strip_imp_prems B |
|
123 |
| strip_imp_prems _ = []; |
|
124 |
||
125 |
(* A1==>...An==>B goes to B, where B is not an implication *) |
|
126 |
fun strip_imp_concl (Const("==>", _) $ A $ B) = strip_imp_concl B |
|
127 |
| strip_imp_concl A = A : term; |
|
128 |
||
129 |
(*Strip and return premises: (i, [], A1==>...Ai==>B) |
|
9460 | 130 |
goes to ([Ai, A(i-1),...,A1] , B) (REVERSED) |
0 | 131 |
if i<0 or else i too big then raises TERM*) |
9460 | 132 |
fun strip_prems (0, As, B) = (As, B) |
133 |
| strip_prems (i, As, Const("==>", _) $ A $ B) = |
|
134 |
strip_prems (i-1, A::As, B) |
|
0 | 135 |
| strip_prems (_, As, A) = raise TERM("strip_prems", A::As); |
136 |
||
16130 | 137 |
(*Count premises -- quicker than (length o strip_prems) *) |
0 | 138 |
fun count_prems (Const("==>", _) $ A $ B, n) = count_prems (B,n+1) |
139 |
| count_prems (_,n) = n; |
|
140 |
||
16130 | 141 |
(*Select Ai from A1 ==>...Ai==>B*) |
142 |
fun nth_prem (1, Const ("==>", _) $ A $ _) = A |
|
143 |
| nth_prem (i, Const ("==>", _) $ _ $ B) = nth_prem (i - 1, B) |
|
144 |
| nth_prem (_, A) = raise TERM ("nth_prem", [A]); |
|
145 |
||
13659
3cf622f6b0b2
Removed obsolete functions dealing with flex-flex constraints.
berghofe
parents:
12902
diff
changeset
|
146 |
(*strip a proof state (Horn clause): |
3cf622f6b0b2
Removed obsolete functions dealing with flex-flex constraints.
berghofe
parents:
12902
diff
changeset
|
147 |
B1 ==> ... Bn ==> C goes to ([B1, ..., Bn], C) *) |
3cf622f6b0b2
Removed obsolete functions dealing with flex-flex constraints.
berghofe
parents:
12902
diff
changeset
|
148 |
fun strip_horn A = (strip_imp_prems A, strip_imp_concl A); |
3cf622f6b0b2
Removed obsolete functions dealing with flex-flex constraints.
berghofe
parents:
12902
diff
changeset
|
149 |
|
4822 | 150 |
|
12137 | 151 |
(** conjunction **) |
152 |
||
19125
59b26248547b
simplified Pure conjunction, based on actual const;
wenzelm
parents:
19103
diff
changeset
|
153 |
val conjunction = Const ("ProtoPure.conjunction", propT --> propT --> propT); |
59b26248547b
simplified Pure conjunction, based on actual const;
wenzelm
parents:
19103
diff
changeset
|
154 |
|
18499 | 155 |
(*A && B*) |
19125
59b26248547b
simplified Pure conjunction, based on actual const;
wenzelm
parents:
19103
diff
changeset
|
156 |
fun mk_conjunction (A, B) = conjunction $ A $ B; |
12137 | 157 |
|
18499 | 158 |
(*A && B && C -- improper*) |
12757 | 159 |
fun mk_conjunction_list [] = Term.all propT $ Abs ("dummy", propT, mk_implies (Bound 0, Bound 0)) |
160 |
| mk_conjunction_list ts = foldr1 mk_conjunction ts; |
|
12137 | 161 |
|
18499 | 162 |
(*(A1 && B1 && C1) && (A2 && B2 && C2 && D2) && A3 && B3 -- improper*) |
163 |
fun mk_conjunction_list2 tss = |
|
164 |
mk_conjunction_list (map mk_conjunction_list (filter_out null tss)); |
|
165 |
||
166 |
(*A && B*) |
|
19125
59b26248547b
simplified Pure conjunction, based on actual const;
wenzelm
parents:
19103
diff
changeset
|
167 |
fun dest_conjunction (Const ("ProtoPure.conjunction", _) $ A $ B) = (A, B) |
18469 | 168 |
| dest_conjunction t = raise TERM ("dest_conjunction", [t]); |
169 |
||
19425 | 170 |
(*A && B && C -- improper*) |
171 |
fun dest_conjunction_list t = |
|
172 |
(case try dest_conjunction t of |
|
173 |
NONE => [t] |
|
174 |
| SOME (A, B) => A :: dest_conjunction_list B); |
|
175 |
||
18499 | 176 |
(*((A && B) && C) && D && E -- flat*) |
18469 | 177 |
fun dest_conjunctions t = |
178 |
(case try dest_conjunction t of |
|
179 |
NONE => [t] |
|
180 |
| SOME (A, B) => dest_conjunctions A @ dest_conjunctions B); |
|
181 |
||
182 |
||
12137 | 183 |
|
398
41f279b477e2
added mk_type, dest_type, mk_inclass, dest_inclass (for axclasses);
wenzelm
parents:
210
diff
changeset
|
184 |
(** types as terms **) |
41f279b477e2
added mk_type, dest_type, mk_inclass, dest_inclass (for axclasses);
wenzelm
parents:
210
diff
changeset
|
185 |
|
19391 | 186 |
fun mk_type ty = Const ("TYPE", Term.itselfT ty); |
398
41f279b477e2
added mk_type, dest_type, mk_inclass, dest_inclass (for axclasses);
wenzelm
parents:
210
diff
changeset
|
187 |
|
41f279b477e2
added mk_type, dest_type, mk_inclass, dest_inclass (for axclasses);
wenzelm
parents:
210
diff
changeset
|
188 |
fun dest_type (Const ("TYPE", Type ("itself", [ty]))) = ty |
41f279b477e2
added mk_type, dest_type, mk_inclass, dest_inclass (for axclasses);
wenzelm
parents:
210
diff
changeset
|
189 |
| dest_type t = raise TERM ("dest_type", [t]); |
41f279b477e2
added mk_type, dest_type, mk_inclass, dest_inclass (for axclasses);
wenzelm
parents:
210
diff
changeset
|
190 |
|
4822 | 191 |
|
19406
410b9d9bf9a1
added mk_classrel, dest_classrel, mk_arities, dest_arity (from axclass.ML);
wenzelm
parents:
19391
diff
changeset
|
192 |
|
410b9d9bf9a1
added mk_classrel, dest_classrel, mk_arities, dest_arity (from axclass.ML);
wenzelm
parents:
19391
diff
changeset
|
193 |
(** type classes **) |
410b9d9bf9a1
added mk_classrel, dest_classrel, mk_arities, dest_arity (from axclass.ML);
wenzelm
parents:
19391
diff
changeset
|
194 |
|
410b9d9bf9a1
added mk_classrel, dest_classrel, mk_arities, dest_arity (from axclass.ML);
wenzelm
parents:
19391
diff
changeset
|
195 |
(* const names *) |
398
41f279b477e2
added mk_type, dest_type, mk_inclass, dest_inclass (for axclasses);
wenzelm
parents:
210
diff
changeset
|
196 |
|
18938 | 197 |
val classN = "_class"; |
198 |
||
199 |
val const_of_class = suffix classN; |
|
19406
410b9d9bf9a1
added mk_classrel, dest_classrel, mk_arities, dest_arity (from axclass.ML);
wenzelm
parents:
19391
diff
changeset
|
200 |
|
18938 | 201 |
fun class_of_const c = unsuffix classN c |
202 |
handle Fail _ => raise TERM ("class_of_const: bad name " ^ quote c, []); |
|
203 |
||
19406
410b9d9bf9a1
added mk_classrel, dest_classrel, mk_arities, dest_arity (from axclass.ML);
wenzelm
parents:
19391
diff
changeset
|
204 |
|
410b9d9bf9a1
added mk_classrel, dest_classrel, mk_arities, dest_arity (from axclass.ML);
wenzelm
parents:
19391
diff
changeset
|
205 |
(* class constraints *) |
410b9d9bf9a1
added mk_classrel, dest_classrel, mk_arities, dest_arity (from axclass.ML);
wenzelm
parents:
19391
diff
changeset
|
206 |
|
398
41f279b477e2
added mk_type, dest_type, mk_inclass, dest_inclass (for axclasses);
wenzelm
parents:
210
diff
changeset
|
207 |
fun mk_inclass (ty, c) = |
19391 | 208 |
Const (const_of_class c, Term.itselfT ty --> propT) $ mk_type ty; |
398
41f279b477e2
added mk_type, dest_type, mk_inclass, dest_inclass (for axclasses);
wenzelm
parents:
210
diff
changeset
|
209 |
|
19406
410b9d9bf9a1
added mk_classrel, dest_classrel, mk_arities, dest_arity (from axclass.ML);
wenzelm
parents:
19391
diff
changeset
|
210 |
fun dest_inclass (t as Const (c_class, _) $ ty) = (dest_type ty, class_of_const c_class) |
398
41f279b477e2
added mk_type, dest_type, mk_inclass, dest_inclass (for axclasses);
wenzelm
parents:
210
diff
changeset
|
211 |
| dest_inclass t = raise TERM ("dest_inclass", [t]); |
41f279b477e2
added mk_type, dest_type, mk_inclass, dest_inclass (for axclasses);
wenzelm
parents:
210
diff
changeset
|
212 |
|
0 | 213 |
|
19406
410b9d9bf9a1
added mk_classrel, dest_classrel, mk_arities, dest_arity (from axclass.ML);
wenzelm
parents:
19391
diff
changeset
|
214 |
(* class relations *) |
410b9d9bf9a1
added mk_classrel, dest_classrel, mk_arities, dest_arity (from axclass.ML);
wenzelm
parents:
19391
diff
changeset
|
215 |
|
410b9d9bf9a1
added mk_classrel, dest_classrel, mk_arities, dest_arity (from axclass.ML);
wenzelm
parents:
19391
diff
changeset
|
216 |
fun mk_classrel (c1, c2) = mk_inclass (Term.aT [c1], c2); |
410b9d9bf9a1
added mk_classrel, dest_classrel, mk_arities, dest_arity (from axclass.ML);
wenzelm
parents:
19391
diff
changeset
|
217 |
|
410b9d9bf9a1
added mk_classrel, dest_classrel, mk_arities, dest_arity (from axclass.ML);
wenzelm
parents:
19391
diff
changeset
|
218 |
fun dest_classrel tm = |
410b9d9bf9a1
added mk_classrel, dest_classrel, mk_arities, dest_arity (from axclass.ML);
wenzelm
parents:
19391
diff
changeset
|
219 |
(case dest_inclass tm of |
410b9d9bf9a1
added mk_classrel, dest_classrel, mk_arities, dest_arity (from axclass.ML);
wenzelm
parents:
19391
diff
changeset
|
220 |
(TVar (_, [c1]), c2) => (c1, c2) |
410b9d9bf9a1
added mk_classrel, dest_classrel, mk_arities, dest_arity (from axclass.ML);
wenzelm
parents:
19391
diff
changeset
|
221 |
| _ => raise TERM ("dest_classrel", [tm])); |
410b9d9bf9a1
added mk_classrel, dest_classrel, mk_arities, dest_arity (from axclass.ML);
wenzelm
parents:
19391
diff
changeset
|
222 |
|
410b9d9bf9a1
added mk_classrel, dest_classrel, mk_arities, dest_arity (from axclass.ML);
wenzelm
parents:
19391
diff
changeset
|
223 |
|
410b9d9bf9a1
added mk_classrel, dest_classrel, mk_arities, dest_arity (from axclass.ML);
wenzelm
parents:
19391
diff
changeset
|
224 |
(* type arities *) |
410b9d9bf9a1
added mk_classrel, dest_classrel, mk_arities, dest_arity (from axclass.ML);
wenzelm
parents:
19391
diff
changeset
|
225 |
|
410b9d9bf9a1
added mk_classrel, dest_classrel, mk_arities, dest_arity (from axclass.ML);
wenzelm
parents:
19391
diff
changeset
|
226 |
fun mk_arities (t, Ss, S) = |
410b9d9bf9a1
added mk_classrel, dest_classrel, mk_arities, dest_arity (from axclass.ML);
wenzelm
parents:
19391
diff
changeset
|
227 |
let val T = Type (t, ListPair.map TFree (Term.invent_names [] "'a" (length Ss), Ss)) |
410b9d9bf9a1
added mk_classrel, dest_classrel, mk_arities, dest_arity (from axclass.ML);
wenzelm
parents:
19391
diff
changeset
|
228 |
in map (fn c => mk_inclass (T, c)) S end; |
410b9d9bf9a1
added mk_classrel, dest_classrel, mk_arities, dest_arity (from axclass.ML);
wenzelm
parents:
19391
diff
changeset
|
229 |
|
410b9d9bf9a1
added mk_classrel, dest_classrel, mk_arities, dest_arity (from axclass.ML);
wenzelm
parents:
19391
diff
changeset
|
230 |
fun dest_arity tm = |
410b9d9bf9a1
added mk_classrel, dest_classrel, mk_arities, dest_arity (from axclass.ML);
wenzelm
parents:
19391
diff
changeset
|
231 |
let |
410b9d9bf9a1
added mk_classrel, dest_classrel, mk_arities, dest_arity (from axclass.ML);
wenzelm
parents:
19391
diff
changeset
|
232 |
fun err () = raise TERM ("dest_arity", [tm]); |
410b9d9bf9a1
added mk_classrel, dest_classrel, mk_arities, dest_arity (from axclass.ML);
wenzelm
parents:
19391
diff
changeset
|
233 |
|
410b9d9bf9a1
added mk_classrel, dest_classrel, mk_arities, dest_arity (from axclass.ML);
wenzelm
parents:
19391
diff
changeset
|
234 |
val (ty, c) = dest_inclass tm; |
410b9d9bf9a1
added mk_classrel, dest_classrel, mk_arities, dest_arity (from axclass.ML);
wenzelm
parents:
19391
diff
changeset
|
235 |
val (t, tvars) = |
410b9d9bf9a1
added mk_classrel, dest_classrel, mk_arities, dest_arity (from axclass.ML);
wenzelm
parents:
19391
diff
changeset
|
236 |
(case ty of |
410b9d9bf9a1
added mk_classrel, dest_classrel, mk_arities, dest_arity (from axclass.ML);
wenzelm
parents:
19391
diff
changeset
|
237 |
Type (t, tys) => (t, map dest_TVar tys handle TYPE _ => err ()) |
410b9d9bf9a1
added mk_classrel, dest_classrel, mk_arities, dest_arity (from axclass.ML);
wenzelm
parents:
19391
diff
changeset
|
238 |
| _ => err ()); |
410b9d9bf9a1
added mk_classrel, dest_classrel, mk_arities, dest_arity (from axclass.ML);
wenzelm
parents:
19391
diff
changeset
|
239 |
val Ss = |
410b9d9bf9a1
added mk_classrel, dest_classrel, mk_arities, dest_arity (from axclass.ML);
wenzelm
parents:
19391
diff
changeset
|
240 |
if has_duplicates (eq_fst (op =)) tvars then err () |
410b9d9bf9a1
added mk_classrel, dest_classrel, mk_arities, dest_arity (from axclass.ML);
wenzelm
parents:
19391
diff
changeset
|
241 |
else map snd tvars; |
410b9d9bf9a1
added mk_classrel, dest_classrel, mk_arities, dest_arity (from axclass.ML);
wenzelm
parents:
19391
diff
changeset
|
242 |
in (t, Ss, c) end; |
410b9d9bf9a1
added mk_classrel, dest_classrel, mk_arities, dest_arity (from axclass.ML);
wenzelm
parents:
19391
diff
changeset
|
243 |
|
410b9d9bf9a1
added mk_classrel, dest_classrel, mk_arities, dest_arity (from axclass.ML);
wenzelm
parents:
19391
diff
changeset
|
244 |
|
18938 | 245 |
|
246 |
(** definitions **) |
|
247 |
||
19103 | 248 |
fun term_kind (Const _) = "existing constant " |
249 |
| term_kind (Free _) = "free variable " |
|
250 |
| term_kind (Bound _) = "bound variable " |
|
251 |
| term_kind _ = ""; |
|
252 |
||
18938 | 253 |
(*c x == t[x] to !!x. c x == t[x]*) |
254 |
fun dest_def pp check_head is_fixed is_fixedT eq = |
|
255 |
let |
|
256 |
fun err msg = raise TERM (msg, [eq]); |
|
19071
fdffd7c40864
dest_def: actually return beta-eta contracted equation;
wenzelm
parents:
18964
diff
changeset
|
257 |
val eq_vars = Term.strip_all_vars eq; |
fdffd7c40864
dest_def: actually return beta-eta contracted equation;
wenzelm
parents:
18964
diff
changeset
|
258 |
val eq_body = Term.strip_all_body eq; |
fdffd7c40864
dest_def: actually return beta-eta contracted equation;
wenzelm
parents:
18964
diff
changeset
|
259 |
|
fdffd7c40864
dest_def: actually return beta-eta contracted equation;
wenzelm
parents:
18964
diff
changeset
|
260 |
val display_terms = commas_quote o map (Pretty.string_of_term pp o Syntax.bound_vars eq_vars); |
18938 | 261 |
val display_types = commas_quote o map (Pretty.string_of_typ pp); |
262 |
||
19071
fdffd7c40864
dest_def: actually return beta-eta contracted equation;
wenzelm
parents:
18964
diff
changeset
|
263 |
val (raw_lhs, rhs) = dest_equals eq_body handle TERM _ => err "Not a meta-equality (==)"; |
fdffd7c40864
dest_def: actually return beta-eta contracted equation;
wenzelm
parents:
18964
diff
changeset
|
264 |
val lhs = Envir.beta_eta_contract raw_lhs; |
fdffd7c40864
dest_def: actually return beta-eta contracted equation;
wenzelm
parents:
18964
diff
changeset
|
265 |
val (head, args) = Term.strip_comb lhs; |
18938 | 266 |
val head_tfrees = Term.add_tfrees head []; |
267 |
||
268 |
fun check_arg (Bound _) = true |
|
269 |
| check_arg (Free (x, _)) = not (is_fixed x) |
|
270 |
| check_arg (Const ("TYPE", Type ("itself", [TFree _]))) = true |
|
271 |
| check_arg _ = false; |
|
272 |
fun close_arg (Bound _) t = t |
|
273 |
| close_arg x t = Term.all (Term.fastype_of x) $ lambda x t; |
|
274 |
||
275 |
val lhs_bads = filter_out check_arg args; |
|
18964 | 276 |
val lhs_dups = duplicates (op aconv) args; |
18938 | 277 |
val rhs_extras = Term.fold_aterms (fn v as Free (x, _) => |
278 |
if is_fixed x orelse member (op aconv) args v then I |
|
279 |
else insert (op aconv) v | _ => I) rhs []; |
|
280 |
val rhs_extrasT = Term.fold_aterms (Term.fold_types (fn v as TFree (a, S) => |
|
281 |
if is_fixedT a orelse member (op =) head_tfrees (a, S) then I |
|
282 |
else insert (op =) v | _ => I)) rhs []; |
|
283 |
in |
|
284 |
if not (check_head head) then |
|
19103 | 285 |
err ("Bad head of lhs: " ^ term_kind head ^ display_terms [head]) |
18938 | 286 |
else if not (null lhs_bads) then |
287 |
err ("Bad arguments on lhs: " ^ display_terms lhs_bads) |
|
288 |
else if not (null lhs_dups) then |
|
289 |
err ("Duplicate arguments on lhs: " ^ display_terms lhs_dups) |
|
290 |
else if not (null rhs_extras) then |
|
291 |
err ("Extra variables on rhs: " ^ display_terms rhs_extras) |
|
292 |
else if not (null rhs_extrasT) then |
|
293 |
err ("Extra type variables on rhs: " ^ display_types rhs_extrasT) |
|
294 |
else if exists_subterm (fn t => t aconv head) rhs then |
|
295 |
err "Entity to be defined occurs on rhs" |
|
19071
fdffd7c40864
dest_def: actually return beta-eta contracted equation;
wenzelm
parents:
18964
diff
changeset
|
296 |
else ((lhs, rhs), fold_rev close_arg args (Term.list_all (eq_vars, (mk_equals (lhs, rhs))))) |
18938 | 297 |
end; |
298 |
||
299 |
(*!!x. c x == t[x] to c == %x. t[x]*) |
|
300 |
fun abs_def eq = |
|
301 |
let |
|
302 |
val body = Term.strip_all_body eq; |
|
303 |
val vars = map Free (Term.rename_wrt_term body (Term.strip_all_vars eq)); |
|
304 |
val (lhs, rhs) = dest_equals (Term.subst_bounds (vars, body)); |
|
305 |
val (lhs', args) = Term.strip_comb lhs; |
|
306 |
val rhs' = Term.list_abs_free (map Term.dest_Free args, rhs); |
|
307 |
in (lhs', rhs') end; |
|
308 |
||
309 |
fun mk_cond_defpair As (lhs, rhs) = |
|
310 |
(case Term.head_of lhs of |
|
311 |
Const (name, _) => |
|
312 |
(NameSpace.base name ^ "_def", list_implies (As, mk_equals (lhs, rhs))) |
|
313 |
| _ => raise TERM ("Malformed definition: head of lhs not a constant", [lhs, rhs])); |
|
314 |
||
315 |
fun mk_defpair lhs_rhs = mk_cond_defpair [] lhs_rhs; |
|
316 |
||
317 |
||
19406
410b9d9bf9a1
added mk_classrel, dest_classrel, mk_arities, dest_arity (from axclass.ML);
wenzelm
parents:
19391
diff
changeset
|
318 |
|
18029
19f1ad18bece
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17120
diff
changeset
|
319 |
(** protected propositions **) |
9460 | 320 |
|
18029
19f1ad18bece
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17120
diff
changeset
|
321 |
val protectC = Const ("prop", propT --> propT); |
19f1ad18bece
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17120
diff
changeset
|
322 |
fun protect t = protectC $ t; |
9460 | 323 |
|
18029
19f1ad18bece
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17120
diff
changeset
|
324 |
fun unprotect (Const ("prop", _) $ t) = t |
19f1ad18bece
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17120
diff
changeset
|
325 |
| unprotect t = raise TERM ("unprotect", [t]); |
9460 | 326 |
|
327 |
||
18181 | 328 |
|
0 | 329 |
(*** Low-level term operations ***) |
330 |
||
331 |
(*Does t occur in u? Or is alpha-convertible to u? |
|
332 |
The term t must contain no loose bound variables*) |
|
16846 | 333 |
fun occs (t, u) = exists_subterm (fn s => t aconv s) u; |
0 | 334 |
|
335 |
(*Close up a formula over all free variables by quantification*) |
|
336 |
fun close_form A = |
|
19425 | 337 |
Term.list_all_free (rev (Term.add_frees A []), A); |
0 | 338 |
|
339 |
||
18938 | 340 |
|
0 | 341 |
(*** Specialized operations for resolution... ***) |
342 |
||
18938 | 343 |
(*computes t(Bound(n+k-1),...,Bound(n)) *) |
344 |
fun combound (t, n, k) = |
|
345 |
if k>0 then combound (t,n+1,k-1) $ (Bound n) else t; |
|
346 |
||
347 |
(* ([xn,...,x1], t) ======> (x1,...,xn)t *) |
|
348 |
fun rlist_abs ([], body) = body |
|
349 |
| rlist_abs ((a,T)::pairs, body) = rlist_abs(pairs, Abs(a, T, body)); |
|
350 |
||
351 |
||
16879
b81d3f2ee565
incr_tvar (from term.ML), incr_indexes: avoid garbage;
wenzelm
parents:
16862
diff
changeset
|
352 |
local exception SAME in |
b81d3f2ee565
incr_tvar (from term.ML), incr_indexes: avoid garbage;
wenzelm
parents:
16862
diff
changeset
|
353 |
|
b81d3f2ee565
incr_tvar (from term.ML), incr_indexes: avoid garbage;
wenzelm
parents:
16862
diff
changeset
|
354 |
fun incrT k = |
b81d3f2ee565
incr_tvar (from term.ML), incr_indexes: avoid garbage;
wenzelm
parents:
16862
diff
changeset
|
355 |
let |
b81d3f2ee565
incr_tvar (from term.ML), incr_indexes: avoid garbage;
wenzelm
parents:
16862
diff
changeset
|
356 |
fun incr (TVar ((a, i), S)) = TVar ((a, i + k), S) |
b81d3f2ee565
incr_tvar (from term.ML), incr_indexes: avoid garbage;
wenzelm
parents:
16862
diff
changeset
|
357 |
| incr (Type (a, Ts)) = Type (a, incrs Ts) |
b81d3f2ee565
incr_tvar (from term.ML), incr_indexes: avoid garbage;
wenzelm
parents:
16862
diff
changeset
|
358 |
| incr _ = raise SAME |
b81d3f2ee565
incr_tvar (from term.ML), incr_indexes: avoid garbage;
wenzelm
parents:
16862
diff
changeset
|
359 |
and incrs (T :: Ts) = |
b81d3f2ee565
incr_tvar (from term.ML), incr_indexes: avoid garbage;
wenzelm
parents:
16862
diff
changeset
|
360 |
(incr T :: (incrs Ts handle SAME => Ts) |
b81d3f2ee565
incr_tvar (from term.ML), incr_indexes: avoid garbage;
wenzelm
parents:
16862
diff
changeset
|
361 |
handle SAME => T :: incrs Ts) |
b81d3f2ee565
incr_tvar (from term.ML), incr_indexes: avoid garbage;
wenzelm
parents:
16862
diff
changeset
|
362 |
| incrs [] = raise SAME; |
b81d3f2ee565
incr_tvar (from term.ML), incr_indexes: avoid garbage;
wenzelm
parents:
16862
diff
changeset
|
363 |
in incr end; |
b81d3f2ee565
incr_tvar (from term.ML), incr_indexes: avoid garbage;
wenzelm
parents:
16862
diff
changeset
|
364 |
|
0 | 365 |
(*For all variables in the term, increment indexnames and lift over the Us |
366 |
result is ?Gidx(B.(lev+n-1),...,B.lev) where lev is abstraction level *) |
|
17120 | 367 |
fun incr_indexes ([], 0) t = t |
18029
19f1ad18bece
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17120
diff
changeset
|
368 |
| incr_indexes (Ts, k) t = |
16879
b81d3f2ee565
incr_tvar (from term.ML), incr_indexes: avoid garbage;
wenzelm
parents:
16862
diff
changeset
|
369 |
let |
18029
19f1ad18bece
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17120
diff
changeset
|
370 |
val n = length Ts; |
16879
b81d3f2ee565
incr_tvar (from term.ML), incr_indexes: avoid garbage;
wenzelm
parents:
16862
diff
changeset
|
371 |
val incrT = if k = 0 then I else incrT k; |
b81d3f2ee565
incr_tvar (from term.ML), incr_indexes: avoid garbage;
wenzelm
parents:
16862
diff
changeset
|
372 |
|
b81d3f2ee565
incr_tvar (from term.ML), incr_indexes: avoid garbage;
wenzelm
parents:
16862
diff
changeset
|
373 |
fun incr lev (Var ((x, i), T)) = |
18938 | 374 |
combound (Var ((x, i + k), Ts ---> (incrT T handle SAME => T)), lev, n) |
16879
b81d3f2ee565
incr_tvar (from term.ML), incr_indexes: avoid garbage;
wenzelm
parents:
16862
diff
changeset
|
375 |
| incr lev (Abs (x, T, body)) = |
b81d3f2ee565
incr_tvar (from term.ML), incr_indexes: avoid garbage;
wenzelm
parents:
16862
diff
changeset
|
376 |
(Abs (x, incrT T, incr (lev + 1) body handle SAME => body) |
b81d3f2ee565
incr_tvar (from term.ML), incr_indexes: avoid garbage;
wenzelm
parents:
16862
diff
changeset
|
377 |
handle SAME => Abs (x, T, incr (lev + 1) body)) |
b81d3f2ee565
incr_tvar (from term.ML), incr_indexes: avoid garbage;
wenzelm
parents:
16862
diff
changeset
|
378 |
| incr lev (t $ u) = |
b81d3f2ee565
incr_tvar (from term.ML), incr_indexes: avoid garbage;
wenzelm
parents:
16862
diff
changeset
|
379 |
(incr lev t $ (incr lev u handle SAME => u) |
b81d3f2ee565
incr_tvar (from term.ML), incr_indexes: avoid garbage;
wenzelm
parents:
16862
diff
changeset
|
380 |
handle SAME => t $ incr lev u) |
b81d3f2ee565
incr_tvar (from term.ML), incr_indexes: avoid garbage;
wenzelm
parents:
16862
diff
changeset
|
381 |
| incr _ (Const (c, T)) = Const (c, incrT T) |
b81d3f2ee565
incr_tvar (from term.ML), incr_indexes: avoid garbage;
wenzelm
parents:
16862
diff
changeset
|
382 |
| incr _ (Free (x, T)) = Free (x, incrT T) |
b81d3f2ee565
incr_tvar (from term.ML), incr_indexes: avoid garbage;
wenzelm
parents:
16862
diff
changeset
|
383 |
| incr _ (t as Bound _) = t; |
b81d3f2ee565
incr_tvar (from term.ML), incr_indexes: avoid garbage;
wenzelm
parents:
16862
diff
changeset
|
384 |
in incr 0 t handle SAME => t end; |
b81d3f2ee565
incr_tvar (from term.ML), incr_indexes: avoid garbage;
wenzelm
parents:
16862
diff
changeset
|
385 |
|
b81d3f2ee565
incr_tvar (from term.ML), incr_indexes: avoid garbage;
wenzelm
parents:
16862
diff
changeset
|
386 |
fun incr_tvar 0 T = T |
b81d3f2ee565
incr_tvar (from term.ML), incr_indexes: avoid garbage;
wenzelm
parents:
16862
diff
changeset
|
387 |
| incr_tvar k T = incrT k T handle SAME => T; |
b81d3f2ee565
incr_tvar (from term.ML), incr_indexes: avoid garbage;
wenzelm
parents:
16862
diff
changeset
|
388 |
|
b81d3f2ee565
incr_tvar (from term.ML), incr_indexes: avoid garbage;
wenzelm
parents:
16862
diff
changeset
|
389 |
end; |
b81d3f2ee565
incr_tvar (from term.ML), incr_indexes: avoid garbage;
wenzelm
parents:
16862
diff
changeset
|
390 |
|
0 | 391 |
|
18248 | 392 |
(* Lifting functions from subgoal and increment: |
18029
19f1ad18bece
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17120
diff
changeset
|
393 |
lift_abs operates on terms |
19f1ad18bece
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17120
diff
changeset
|
394 |
lift_all operates on propositions *) |
19f1ad18bece
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17120
diff
changeset
|
395 |
|
19f1ad18bece
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17120
diff
changeset
|
396 |
fun lift_abs inc = |
19f1ad18bece
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17120
diff
changeset
|
397 |
let |
19f1ad18bece
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17120
diff
changeset
|
398 |
fun lift Ts (Const ("==>", _) $ _ $ B) t = lift Ts B t |
18248 | 399 |
| lift Ts (Const ("all", _) $ Abs (a, T, B)) t = Abs (a, T, lift (T :: Ts) B t) |
18029
19f1ad18bece
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17120
diff
changeset
|
400 |
| lift Ts _ t = incr_indexes (rev Ts, inc) t; |
19f1ad18bece
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17120
diff
changeset
|
401 |
in lift [] end; |
19f1ad18bece
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17120
diff
changeset
|
402 |
|
19f1ad18bece
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17120
diff
changeset
|
403 |
fun lift_all inc = |
19f1ad18bece
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17120
diff
changeset
|
404 |
let |
19f1ad18bece
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17120
diff
changeset
|
405 |
fun lift Ts ((c as Const ("==>", _)) $ A $ B) t = c $ A $ lift Ts B t |
18248 | 406 |
| lift Ts ((c as Const ("all", _)) $ Abs (a, T, B)) t = c $ Abs (a, T, lift (T :: Ts) B t) |
18029
19f1ad18bece
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17120
diff
changeset
|
407 |
| lift Ts _ t = incr_indexes (rev Ts, inc) t; |
19f1ad18bece
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17120
diff
changeset
|
408 |
in lift [] end; |
0 | 409 |
|
410 |
(*Strips assumptions in goal, yielding list of hypotheses. *) |
|
411 |
fun strip_assums_hyp (Const("==>", _) $ H $ B) = H :: strip_assums_hyp B |
|
412 |
| strip_assums_hyp (Const("all",_)$Abs(a,T,t)) = strip_assums_hyp t |
|
413 |
| strip_assums_hyp B = []; |
|
414 |
||
415 |
(*Strips assumptions in goal, yielding conclusion. *) |
|
416 |
fun strip_assums_concl (Const("==>", _) $ H $ B) = strip_assums_concl B |
|
417 |
| strip_assums_concl (Const("all",_)$Abs(a,T,t)) = strip_assums_concl t |
|
418 |
| strip_assums_concl B = B; |
|
419 |
||
420 |
(*Make a list of all the parameters in a subgoal, even if nested*) |
|
421 |
fun strip_params (Const("==>", _) $ H $ B) = strip_params B |
|
422 |
| strip_params (Const("all",_)$Abs(a,T,t)) = (a,T) :: strip_params t |
|
423 |
| strip_params B = []; |
|
424 |
||
9667 | 425 |
(*test for meta connectives in prems of a 'subgoal'*) |
426 |
fun has_meta_prems prop i = |
|
427 |
let |
|
428 |
fun is_meta (Const ("==>", _) $ _ $ _) = true |
|
10442 | 429 |
| is_meta (Const ("==", _) $ _ $ _) = true |
9667 | 430 |
| is_meta (Const ("all", _) $ _) = true |
431 |
| is_meta _ = false; |
|
432 |
in |
|
13659
3cf622f6b0b2
Removed obsolete functions dealing with flex-flex constraints.
berghofe
parents:
12902
diff
changeset
|
433 |
(case strip_prems (i, [], prop) of |
9667 | 434 |
(B :: _, _) => exists is_meta (strip_assums_hyp B) |
435 |
| _ => false) handle TERM _ => false |
|
436 |
end; |
|
9483 | 437 |
|
0 | 438 |
(*Removes the parameters from a subgoal and renumber bvars in hypotheses, |
9460 | 439 |
where j is the total number of parameters (precomputed) |
0 | 440 |
If n>0 then deletes assumption n. *) |
9460 | 441 |
fun remove_params j n A = |
0 | 442 |
if j=0 andalso n<=0 then A (*nothing left to do...*) |
443 |
else case A of |
|
9460 | 444 |
Const("==>", _) $ H $ B => |
445 |
if n=1 then (remove_params j (n-1) B) |
|
446 |
else implies $ (incr_boundvars j H) $ (remove_params j (n-1) B) |
|
0 | 447 |
| Const("all",_)$Abs(a,T,t) => remove_params (j-1) n t |
448 |
| _ => if n>0 then raise TERM("remove_params", [A]) |
|
449 |
else A; |
|
450 |
||
451 |
(** Auto-renaming of parameters in subgoals **) |
|
452 |
||
453 |
val auto_rename = ref false |
|
454 |
and rename_prefix = ref "ka"; |
|
455 |
||
456 |
(*rename_prefix is not exported; it is set by this function.*) |
|
457 |
fun set_rename_prefix a = |
|
4693 | 458 |
if a<>"" andalso forall Symbol.is_letter (Symbol.explode a) |
0 | 459 |
then (rename_prefix := a; auto_rename := true) |
460 |
else error"rename prefix must be nonempty and consist of letters"; |
|
461 |
||
462 |
(*Makes parameters in a goal have distinctive names (not guaranteed unique!) |
|
463 |
A name clash could cause the printer to rename bound vars; |
|
464 |
then res_inst_tac would not work properly.*) |
|
465 |
fun rename_vars (a, []) = [] |
|
466 |
| rename_vars (a, (_,T)::vars) = |
|
12902 | 467 |
(a,T) :: rename_vars (Symbol.bump_string a, vars); |
0 | 468 |
|
469 |
(*Move all parameters to the front of the subgoal, renaming them apart; |
|
470 |
if n>0 then deletes assumption n. *) |
|
471 |
fun flatten_params n A = |
|
472 |
let val params = strip_params A; |
|
9460 | 473 |
val vars = if !auto_rename |
474 |
then rename_vars (!rename_prefix, params) |
|
475 |
else ListPair.zip (variantlist(map #1 params,[]), |
|
476 |
map #2 params) |
|
0 | 477 |
in list_all (vars, remove_params (length vars) n A) |
478 |
end; |
|
479 |
||
480 |
(*Makes parameters in a goal have the names supplied by the list cs.*) |
|
481 |
fun list_rename_params (cs, Const("==>", _) $ A $ B) = |
|
482 |
implies $ A $ list_rename_params (cs, B) |
|
9460 | 483 |
| list_rename_params (c::cs, Const("all",_)$Abs(_,T,t)) = |
0 | 484 |
all T $ Abs(c, T, list_rename_params (cs, t)) |
485 |
| list_rename_params (cs, B) = B; |
|
486 |
||
15451
c6c8786b9921
fixed thin_tac with higher-level assumptions by removing the old code to
paulson
parents:
14137
diff
changeset
|
487 |
(*** Treatmsent of "assume", "erule", etc. ***) |
0 | 488 |
|
16879
b81d3f2ee565
incr_tvar (from term.ML), incr_indexes: avoid garbage;
wenzelm
parents:
16862
diff
changeset
|
489 |
(*Strips assumptions in goal yielding |
15451
c6c8786b9921
fixed thin_tac with higher-level assumptions by removing the old code to
paulson
parents:
14137
diff
changeset
|
490 |
HS = [Hn,...,H1], params = [xm,...,x1], and B, |
16879
b81d3f2ee565
incr_tvar (from term.ML), incr_indexes: avoid garbage;
wenzelm
parents:
16862
diff
changeset
|
491 |
where x1...xm are the parameters. This version (21.1.2005) REQUIRES |
b81d3f2ee565
incr_tvar (from term.ML), incr_indexes: avoid garbage;
wenzelm
parents:
16862
diff
changeset
|
492 |
the the parameters to be flattened, but it allows erule to work on |
15451
c6c8786b9921
fixed thin_tac with higher-level assumptions by removing the old code to
paulson
parents:
14137
diff
changeset
|
493 |
assumptions of the form !!x. phi. Any !! after the outermost string |
c6c8786b9921
fixed thin_tac with higher-level assumptions by removing the old code to
paulson
parents:
14137
diff
changeset
|
494 |
will be regarded as belonging to the conclusion, and left untouched. |
15454 | 495 |
Used ONLY by assum_pairs. |
496 |
Unless nasms<0, it can terminate the recursion early; that allows |
|
497 |
erule to work on assumptions of the form P==>Q.*) |
|
498 |
fun strip_assums_imp (0, Hs, B) = (Hs, B) (*recursion terminated by nasms*) |
|
16879
b81d3f2ee565
incr_tvar (from term.ML), incr_indexes: avoid garbage;
wenzelm
parents:
16862
diff
changeset
|
499 |
| strip_assums_imp (nasms, Hs, Const("==>", _) $ H $ B) = |
15454 | 500 |
strip_assums_imp (nasms-1, H::Hs, B) |
501 |
| strip_assums_imp (_, Hs, B) = (Hs, B); (*recursion terminated by B*) |
|
502 |
||
0 | 503 |
|
15451
c6c8786b9921
fixed thin_tac with higher-level assumptions by removing the old code to
paulson
parents:
14137
diff
changeset
|
504 |
(*Strips OUTER parameters only, unlike similar legacy versions.*) |
c6c8786b9921
fixed thin_tac with higher-level assumptions by removing the old code to
paulson
parents:
14137
diff
changeset
|
505 |
fun strip_assums_all (params, Const("all",_)$Abs(a,T,t)) = |
c6c8786b9921
fixed thin_tac with higher-level assumptions by removing the old code to
paulson
parents:
14137
diff
changeset
|
506 |
strip_assums_all ((a,T)::params, t) |
c6c8786b9921
fixed thin_tac with higher-level assumptions by removing the old code to
paulson
parents:
14137
diff
changeset
|
507 |
| strip_assums_all (params, B) = (params, B); |
0 | 508 |
|
509 |
(*Produces disagreement pairs, one for each assumption proof, in order. |
|
510 |
A is the first premise of the lifted rule, and thus has the form |
|
15454 | 511 |
H1 ==> ... Hk ==> B and the pairs are (H1,B),...,(Hk,B). |
512 |
nasms is the number of assumptions in the original subgoal, needed when B |
|
513 |
has the form B1 ==> B2: it stops B1 from being taken as an assumption. *) |
|
514 |
fun assum_pairs(nasms,A) = |
|
15451
c6c8786b9921
fixed thin_tac with higher-level assumptions by removing the old code to
paulson
parents:
14137
diff
changeset
|
515 |
let val (params, A') = strip_assums_all ([],A) |
15454 | 516 |
val (Hs,B) = strip_assums_imp (nasms,[],A') |
18938 | 517 |
fun abspar t = rlist_abs(params, t) |
15451
c6c8786b9921
fixed thin_tac with higher-level assumptions by removing the old code to
paulson
parents:
14137
diff
changeset
|
518 |
val D = abspar B |
c6c8786b9921
fixed thin_tac with higher-level assumptions by removing the old code to
paulson
parents:
14137
diff
changeset
|
519 |
fun pairrev ([], pairs) = pairs |
c6c8786b9921
fixed thin_tac with higher-level assumptions by removing the old code to
paulson
parents:
14137
diff
changeset
|
520 |
| pairrev (H::Hs, pairs) = pairrev(Hs, (abspar H, D) :: pairs) |
c6c8786b9921
fixed thin_tac with higher-level assumptions by removing the old code to
paulson
parents:
14137
diff
changeset
|
521 |
in pairrev (Hs,[]) |
0 | 522 |
end; |
523 |
||
18938 | 524 |
(*Converts Frees to Vars and TFrees to TVars*) |
16862 | 525 |
fun varify (Const(a, T)) = Const (a, Type.varifyT T) |
526 |
| varify (Free (a, T)) = Var ((a, 0), Type.varifyT T) |
|
527 |
| varify (Var (ixn, T)) = Var (ixn, Type.varifyT T) |
|
528 |
| varify (t as Bound _) = t |
|
529 |
| varify (Abs (a, T, body)) = Abs (a, Type.varifyT T, varify body) |
|
530 |
| varify (f $ t) = varify f $ varify t; |
|
0 | 531 |
|
18938 | 532 |
(*Inverse of varify.*) |
16862 | 533 |
fun unvarify (Const (a, T)) = Const (a, Type.unvarifyT T) |
534 |
| unvarify (Free (a, T)) = Free (a, Type.unvarifyT T) |
|
535 |
| unvarify (Var ((a, 0), T)) = Free (a, Type.unvarifyT T) |
|
536 |
| unvarify (Var (ixn, T)) = Var (ixn, Type.unvarifyT T) (*non-0 index!*) |
|
537 |
| unvarify (t as Bound _) = t |
|
538 |
| unvarify (Abs (a, T, body)) = Abs (a, Type.unvarifyT T, unvarify body) |
|
539 |
| unvarify (f $ t) = unvarify f $ unvarify t; |
|
546 | 540 |
|
13799
77614fe09362
Moved get_goal, prems_of_goal and concl_of_goal from goals.ML to logic.ML
berghofe
parents:
13659
diff
changeset
|
541 |
|
16862 | 542 |
(* goal states *) |
543 |
||
544 |
fun get_goal st i = nth_prem (i, st) |
|
545 |
handle TERM _ => error "Goal number out of range"; |
|
13799
77614fe09362
Moved get_goal, prems_of_goal and concl_of_goal from goals.ML to logic.ML
berghofe
parents:
13659
diff
changeset
|
546 |
|
77614fe09362
Moved get_goal, prems_of_goal and concl_of_goal from goals.ML to logic.ML
berghofe
parents:
13659
diff
changeset
|
547 |
(*reverses parameters for substitution*) |
77614fe09362
Moved get_goal, prems_of_goal and concl_of_goal from goals.ML to logic.ML
berghofe
parents:
13659
diff
changeset
|
548 |
fun goal_params st i = |
77614fe09362
Moved get_goal, prems_of_goal and concl_of_goal from goals.ML to logic.ML
berghofe
parents:
13659
diff
changeset
|
549 |
let val gi = get_goal st i |
14137
c57ec95e7763
Removed extraneous rev in function goal_params (the list of parameters
berghofe
parents:
14107
diff
changeset
|
550 |
val rfrees = map Free (rename_wrt_term gi (strip_params gi)) |
13799
77614fe09362
Moved get_goal, prems_of_goal and concl_of_goal from goals.ML to logic.ML
berghofe
parents:
13659
diff
changeset
|
551 |
in (gi, rfrees) end; |
77614fe09362
Moved get_goal, prems_of_goal and concl_of_goal from goals.ML to logic.ML
berghofe
parents:
13659
diff
changeset
|
552 |
|
77614fe09362
Moved get_goal, prems_of_goal and concl_of_goal from goals.ML to logic.ML
berghofe
parents:
13659
diff
changeset
|
553 |
fun concl_of_goal st i = |
77614fe09362
Moved get_goal, prems_of_goal and concl_of_goal from goals.ML to logic.ML
berghofe
parents:
13659
diff
changeset
|
554 |
let val (gi, rfrees) = goal_params st i |
77614fe09362
Moved get_goal, prems_of_goal and concl_of_goal from goals.ML to logic.ML
berghofe
parents:
13659
diff
changeset
|
555 |
val B = strip_assums_concl gi |
77614fe09362
Moved get_goal, prems_of_goal and concl_of_goal from goals.ML to logic.ML
berghofe
parents:
13659
diff
changeset
|
556 |
in subst_bounds (rfrees, B) end; |
77614fe09362
Moved get_goal, prems_of_goal and concl_of_goal from goals.ML to logic.ML
berghofe
parents:
13659
diff
changeset
|
557 |
|
77614fe09362
Moved get_goal, prems_of_goal and concl_of_goal from goals.ML to logic.ML
berghofe
parents:
13659
diff
changeset
|
558 |
fun prems_of_goal st i = |
77614fe09362
Moved get_goal, prems_of_goal and concl_of_goal from goals.ML to logic.ML
berghofe
parents:
13659
diff
changeset
|
559 |
let val (gi, rfrees) = goal_params st i |
77614fe09362
Moved get_goal, prems_of_goal and concl_of_goal from goals.ML to logic.ML
berghofe
parents:
13659
diff
changeset
|
560 |
val As = strip_assums_hyp gi |
77614fe09362
Moved get_goal, prems_of_goal and concl_of_goal from goals.ML to logic.ML
berghofe
parents:
13659
diff
changeset
|
561 |
in map (fn A => subst_bounds (rfrees, A)) As end; |
77614fe09362
Moved get_goal, prems_of_goal and concl_of_goal from goals.ML to logic.ML
berghofe
parents:
13659
diff
changeset
|
562 |
|
0 | 563 |
end; |