| author | wenzelm | 
| Fri, 15 Oct 2021 19:25:31 +0200 | |
| changeset 74525 | c960bfcb91db | 
| parent 74518 | de4f151c2a34 | 
| child 74532 | 64d1b02327a4 | 
| permissions | -rw-r--r-- | 
| 62058 | 1  | 
(* Title: HOL/Library/old_recdef.ML  | 
| 60520 | 2  | 
Author: Konrad Slind, Cambridge University Computer Laboratory  | 
3  | 
Author: Lucas Dixon, University of Edinburgh  | 
|
4  | 
||
5  | 
Old TFL/recdef package.  | 
|
6  | 
*)  | 
|
7  | 
||
8  | 
signature CASE_SPLIT =  | 
|
9  | 
sig  | 
|
10  | 
(* try to recursively split conjectured thm to given list of thms *)  | 
|
11  | 
val splitto : Proof.context -> thm list -> thm -> thm  | 
|
12  | 
end;  | 
|
13  | 
||
14  | 
signature UTILS =  | 
|
15  | 
sig  | 
|
16  | 
  exception ERR of {module: string, func: string, mesg: string}
 | 
|
17  | 
  val end_itlist: ('a -> 'a -> 'a) -> 'a list -> 'a
 | 
|
18  | 
  val itlist2: ('a -> 'b -> 'c -> 'c) -> 'a list -> 'b list -> 'c -> 'c
 | 
|
19  | 
  val pluck: ('a -> bool) -> 'a list -> 'a * 'a list
 | 
|
20  | 
  val zip3: 'a list -> 'b list -> 'c list -> ('a*'b*'c) list
 | 
|
21  | 
  val take: ('a -> 'b) -> int * 'a list -> 'b list
 | 
|
22  | 
end;  | 
|
23  | 
||
24  | 
signature USYNTAX =  | 
|
25  | 
sig  | 
|
26  | 
  datatype lambda = VAR   of {Name : string, Ty : typ}
 | 
|
27  | 
                  | CONST of {Name : string, Ty : typ}
 | 
|
28  | 
                  | COMB  of {Rator: term, Rand : term}
 | 
|
29  | 
                  | LAMB  of {Bvar : term, Body : term}
 | 
|
30  | 
||
31  | 
val alpha : typ  | 
|
32  | 
||
33  | 
(* Types *)  | 
|
34  | 
val type_vars : typ -> typ list  | 
|
35  | 
val type_varsl : typ list -> typ list  | 
|
36  | 
val mk_vartype : string -> typ  | 
|
37  | 
val is_vartype : typ -> bool  | 
|
38  | 
val strip_prod_type : typ -> typ list  | 
|
39  | 
||
40  | 
(* Terms *)  | 
|
41  | 
val free_vars_lr : term -> term list  | 
|
42  | 
val type_vars_in_term : term -> typ list  | 
|
43  | 
val dest_term : term -> lambda  | 
|
44  | 
||
45  | 
(* Prelogic *)  | 
|
46  | 
val inst : (typ*typ) list -> term -> term  | 
|
47  | 
||
48  | 
(* Construction routines *)  | 
|
49  | 
  val mk_abs    :{Bvar  : term, Body : term} -> term
 | 
|
50  | 
||
51  | 
  val mk_imp    :{ant : term, conseq :  term} -> term
 | 
|
52  | 
  val mk_select :{Bvar : term, Body : term} -> term
 | 
|
53  | 
  val mk_forall :{Bvar : term, Body : term} -> term
 | 
|
54  | 
  val mk_exists :{Bvar : term, Body : term} -> term
 | 
|
55  | 
  val mk_conj   :{conj1 : term, conj2 : term} -> term
 | 
|
56  | 
  val mk_disj   :{disj1 : term, disj2 : term} -> term
 | 
|
57  | 
  val mk_pabs   :{varstruct : term, body : term} -> term
 | 
|
58  | 
||
59  | 
(* Destruction routines *)  | 
|
60  | 
  val dest_const: term -> {Name : string, Ty : typ}
 | 
|
61  | 
  val dest_comb : term -> {Rator : term, Rand : term}
 | 
|
62  | 
  val dest_abs  : string list -> term -> {Bvar : term, Body : term} * string list
 | 
|
63  | 
  val dest_eq     : term -> {lhs : term, rhs : term}
 | 
|
64  | 
  val dest_imp    : term -> {ant : term, conseq : term}
 | 
|
65  | 
  val dest_forall : term -> {Bvar : term, Body : term}
 | 
|
66  | 
  val dest_exists : term -> {Bvar : term, Body : term}
 | 
|
67  | 
val dest_neg : term -> term  | 
|
68  | 
  val dest_conj   : term -> {conj1 : term, conj2 : term}
 | 
|
69  | 
  val dest_disj   : term -> {disj1 : term, disj2 : term}
 | 
|
70  | 
  val dest_pair   : term -> {fst : term, snd : term}
 | 
|
71  | 
  val dest_pabs   : string list -> term -> {varstruct : term, body : term, used : string list}
 | 
|
72  | 
||
73  | 
val lhs : term -> term  | 
|
74  | 
val rhs : term -> term  | 
|
75  | 
val rand : term -> term  | 
|
76  | 
||
77  | 
(* Query routines *)  | 
|
78  | 
val is_imp : term -> bool  | 
|
79  | 
val is_forall : term -> bool  | 
|
80  | 
val is_exists : term -> bool  | 
|
81  | 
val is_neg : term -> bool  | 
|
82  | 
val is_conj : term -> bool  | 
|
83  | 
val is_disj : term -> bool  | 
|
84  | 
val is_pair : term -> bool  | 
|
85  | 
val is_pabs : term -> bool  | 
|
86  | 
||
87  | 
(* Construction of a term from a list of Preterms *)  | 
|
88  | 
val list_mk_abs : (term list * term) -> term  | 
|
89  | 
val list_mk_imp : (term list * term) -> term  | 
|
90  | 
val list_mk_forall : (term list * term) -> term  | 
|
91  | 
val list_mk_conj : term list -> term  | 
|
92  | 
||
93  | 
(* Destructing a term to a list of Preterms *)  | 
|
94  | 
val strip_comb : term -> (term * term list)  | 
|
95  | 
val strip_abs : term -> (term list * term)  | 
|
96  | 
val strip_imp : term -> (term list * term)  | 
|
97  | 
val strip_forall : term -> (term list * term)  | 
|
98  | 
val strip_exists : term -> (term list * term)  | 
|
99  | 
val strip_disj : term -> term list  | 
|
100  | 
||
101  | 
(* Miscellaneous *)  | 
|
102  | 
val mk_vstruct : typ -> term list -> term  | 
|
103  | 
val gen_all : term -> term  | 
|
104  | 
val find_term : (term -> bool) -> term -> term option  | 
|
105  | 
val dest_relation : term -> term * term * term  | 
|
106  | 
val is_WFR : term -> bool  | 
|
107  | 
val ARB : typ -> term  | 
|
108  | 
end;  | 
|
109  | 
||
110  | 
signature DCTERM =  | 
|
111  | 
sig  | 
|
112  | 
val dest_comb: cterm -> cterm * cterm  | 
|
| 74518 | 113  | 
val dest_abs: cterm -> cterm * cterm  | 
| 60520 | 114  | 
val capply: cterm -> cterm -> cterm  | 
115  | 
val cabs: cterm -> cterm -> cterm  | 
|
116  | 
val mk_conj: cterm * cterm -> cterm  | 
|
117  | 
val mk_disj: cterm * cterm -> cterm  | 
|
118  | 
val mk_exists: cterm * cterm -> cterm  | 
|
119  | 
val dest_conj: cterm -> cterm * cterm  | 
|
120  | 
  val dest_const: cterm -> {Name: string, Ty: typ}
 | 
|
121  | 
val dest_disj: cterm -> cterm * cterm  | 
|
122  | 
val dest_eq: cterm -> cterm * cterm  | 
|
123  | 
val dest_exists: cterm -> cterm * cterm  | 
|
124  | 
val dest_forall: cterm -> cterm * cterm  | 
|
125  | 
val dest_imp: cterm -> cterm * cterm  | 
|
126  | 
val dest_neg: cterm -> cterm  | 
|
127  | 
val dest_pair: cterm -> cterm * cterm  | 
|
128  | 
  val dest_var: cterm -> {Name:string, Ty:typ}
 | 
|
129  | 
val is_conj: cterm -> bool  | 
|
130  | 
val is_disj: cterm -> bool  | 
|
131  | 
val is_eq: cterm -> bool  | 
|
132  | 
val is_exists: cterm -> bool  | 
|
133  | 
val is_forall: cterm -> bool  | 
|
134  | 
val is_imp: cterm -> bool  | 
|
135  | 
val is_neg: cterm -> bool  | 
|
136  | 
val is_pair: cterm -> bool  | 
|
137  | 
val list_mk_disj: cterm list -> cterm  | 
|
138  | 
val strip_abs: cterm -> cterm list * cterm  | 
|
139  | 
val strip_comb: cterm -> cterm * cterm list  | 
|
140  | 
val strip_disj: cterm -> cterm list  | 
|
141  | 
val strip_exists: cterm -> cterm list * cterm  | 
|
142  | 
val strip_forall: cterm -> cterm list * cterm  | 
|
143  | 
val strip_imp: cterm -> cterm list * cterm  | 
|
144  | 
val drop_prop: cterm -> cterm  | 
|
145  | 
val mk_prop: cterm -> cterm  | 
|
146  | 
end;  | 
|
147  | 
||
148  | 
signature RULES =  | 
|
149  | 
sig  | 
|
150  | 
val dest_thm: thm -> term list * term  | 
|
151  | 
||
152  | 
(* Inference rules *)  | 
|
153  | 
val REFL: cterm -> thm  | 
|
154  | 
val ASSUME: cterm -> thm  | 
|
155  | 
val MP: thm -> thm -> thm  | 
|
156  | 
val MATCH_MP: thm -> thm -> thm  | 
|
157  | 
val CONJUNCT1: thm -> thm  | 
|
158  | 
val CONJUNCT2: thm -> thm  | 
|
159  | 
val CONJUNCTS: thm -> thm list  | 
|
160  | 
val DISCH: cterm -> thm -> thm  | 
|
161  | 
val UNDISCH: thm -> thm  | 
|
162  | 
val SPEC: cterm -> thm -> thm  | 
|
163  | 
val ISPEC: cterm -> thm -> thm  | 
|
164  | 
val ISPECL: cterm list -> thm -> thm  | 
|
165  | 
val GEN: Proof.context -> cterm -> thm -> thm  | 
|
166  | 
val GENL: Proof.context -> cterm list -> thm -> thm  | 
|
167  | 
val LIST_CONJ: thm list -> thm  | 
|
168  | 
||
169  | 
val SYM: thm -> thm  | 
|
170  | 
val DISCH_ALL: thm -> thm  | 
|
171  | 
val FILTER_DISCH_ALL: (term -> bool) -> thm -> thm  | 
|
172  | 
val SPEC_ALL: thm -> thm  | 
|
173  | 
val GEN_ALL: Proof.context -> thm -> thm  | 
|
174  | 
val IMP_TRANS: thm -> thm -> thm  | 
|
175  | 
val PROVE_HYP: thm -> thm -> thm  | 
|
176  | 
||
177  | 
val CHOOSE: Proof.context -> cterm * thm -> thm -> thm  | 
|
| 60781 | 178  | 
val EXISTS: Proof.context -> cterm * cterm -> thm -> thm  | 
| 60520 | 179  | 
val IT_EXISTS: Proof.context -> (cterm * cterm) list -> thm -> thm  | 
180  | 
||
181  | 
val EVEN_ORS: thm list -> thm list  | 
|
182  | 
val DISJ_CASESL: thm -> thm list -> thm  | 
|
183  | 
||
184  | 
val list_beta_conv: cterm -> cterm list -> thm  | 
|
185  | 
val SUBS: Proof.context -> thm list -> thm -> thm  | 
|
186  | 
val simpl_conv: Proof.context -> thm list -> cterm -> thm  | 
|
187  | 
||
188  | 
val rbeta: thm -> thm  | 
|
189  | 
val tracing: bool Unsynchronized.ref  | 
|
190  | 
val CONTEXT_REWRITE_RULE: Proof.context ->  | 
|
191  | 
term * term list * thm * thm list -> thm -> thm * term list  | 
|
192  | 
val RIGHT_ASSOC: Proof.context -> thm -> thm  | 
|
193  | 
||
194  | 
val prove: Proof.context -> bool -> term * tactic -> thm  | 
|
195  | 
end;  | 
|
196  | 
||
197  | 
signature THRY =  | 
|
198  | 
sig  | 
|
199  | 
val match_term: theory -> term -> term -> (term * term) list * (typ * typ) list  | 
|
200  | 
val match_type: theory -> typ -> typ -> (typ * typ) list  | 
|
201  | 
val typecheck: theory -> term -> cterm  | 
|
202  | 
(*datatype facts of various flavours*)  | 
|
203  | 
  val match_info: theory -> string -> {constructors: term list, case_const: term} option
 | 
|
204  | 
  val induct_info: theory -> string -> {constructors: term list, nchotomy: thm} option
 | 
|
205  | 
  val extract_info: theory -> {case_congs: thm list, case_rewrites: thm list}
 | 
|
206  | 
end;  | 
|
207  | 
||
208  | 
signature PRIM =  | 
|
209  | 
sig  | 
|
210  | 
val trace: bool Unsynchronized.ref  | 
|
211  | 
val trace_thms: Proof.context -> string -> thm list -> unit  | 
|
212  | 
val trace_cterm: Proof.context -> string -> cterm -> unit  | 
|
213  | 
type pattern  | 
|
214  | 
  val mk_functional: theory -> term list -> {functional: term, pats: pattern list}
 | 
|
215  | 
val wfrec_definition0: string -> term -> term -> theory -> thm * theory  | 
|
216  | 
val post_definition: Proof.context -> thm list -> thm * pattern list ->  | 
|
217  | 
   {rules: thm,
 | 
|
218  | 
rows: int list,  | 
|
219  | 
TCs: term list list,  | 
|
220  | 
full_pats_TCs: (term * term list) list}  | 
|
| 60699 | 221  | 
val mk_induction: Proof.context ->  | 
| 60520 | 222  | 
    {fconst: term, R: term, SV: term list, pat_TCs_list: (term * term list) list} -> thm
 | 
223  | 
val postprocess: Proof.context -> bool ->  | 
|
224  | 
    {wf_tac: tactic, terminator: tactic, simplifier: cterm -> thm} ->
 | 
|
225  | 
    {rules: thm, induction: thm, TCs: term list list} ->
 | 
|
226  | 
    {rules: thm, induction: thm, nested_tcs: thm list}
 | 
|
227  | 
end;  | 
|
228  | 
||
229  | 
signature TFL =  | 
|
230  | 
sig  | 
|
231  | 
val define_i: bool -> thm list -> thm list -> xstring -> term -> term list -> Proof.context ->  | 
|
232  | 
    {lhs: term, rules: (thm * int) list, induct: thm, tcs: term list} * Proof.context
 | 
|
233  | 
val define: bool -> thm list -> thm list -> xstring -> string -> string list -> Proof.context ->  | 
|
234  | 
    {lhs: term, rules: (thm * int) list, induct: thm, tcs: term list} * Proof.context
 | 
|
235  | 
end;  | 
|
236  | 
||
237  | 
signature OLD_RECDEF =  | 
|
238  | 
sig  | 
|
239  | 
val get_recdef: theory -> string  | 
|
240  | 
    -> {lhs: term, simps: thm list, rules: thm list list, induct: thm, tcs: term list} option
 | 
|
241  | 
  val get_hints: Proof.context -> {simps: thm list, congs: (string * thm) list, wfs: thm list}
 | 
|
242  | 
val simp_add: attribute  | 
|
243  | 
val simp_del: attribute  | 
|
244  | 
val cong_add: attribute  | 
|
245  | 
val cong_del: attribute  | 
|
246  | 
val wf_add: attribute  | 
|
247  | 
val wf_del: attribute  | 
|
248  | 
val add_recdef: bool -> xstring -> string -> ((binding * string) * Token.src list) list ->  | 
|
249  | 
Token.src option -> theory -> theory  | 
|
250  | 
      * {lhs: term, simps: thm list, rules: thm list list, induct: thm, tcs: term list}
 | 
|
251  | 
val add_recdef_i: bool -> xstring -> term -> ((binding * term) * attribute list) list ->  | 
|
252  | 
    theory -> theory * {lhs: term, simps: thm list, rules: thm list list, induct: thm, tcs: term list}
 | 
|
253  | 
end;  | 
|
254  | 
||
255  | 
structure Old_Recdef: OLD_RECDEF =  | 
|
256  | 
struct  | 
|
257  | 
||
258  | 
(*** extra case splitting for TFL ***)  | 
|
259  | 
||
260  | 
structure CaseSplit: CASE_SPLIT =  | 
|
261  | 
struct  | 
|
262  | 
||
263  | 
(* make a casethm from an induction thm *)  | 
|
| 60752 | 264  | 
fun cases_thm_of_induct_thm ctxt =  | 
265  | 
Seq.hd o (ALLGOALS (fn i => REPEAT (eresolve_tac ctxt [Drule.thin_rl] i)));  | 
|
| 60520 | 266  | 
|
267  | 
(* get the case_thm (my version) from a type *)  | 
|
| 60752 | 268  | 
fun case_thm_of_ty ctxt ty =  | 
| 60520 | 269  | 
let  | 
| 60752 | 270  | 
val thy = Proof_Context.theory_of ctxt  | 
| 60520 | 271  | 
val ty_str = case ty of  | 
272  | 
Type(ty_str, _) => ty_str  | 
|
273  | 
                   | TFree(s,_)  => error ("Free type: " ^ s)
 | 
|
| 60521 | 274  | 
                   | TVar((s,_),_) => error ("Free variable: " ^ s)
 | 
| 60520 | 275  | 
      val {induct, ...} = BNF_LFP_Compat.the_info thy [BNF_LFP_Compat.Keep_Nesting] ty_str
 | 
276  | 
in  | 
|
| 60752 | 277  | 
cases_thm_of_induct_thm ctxt induct  | 
| 60520 | 278  | 
end;  | 
279  | 
||
280  | 
||
281  | 
(* for use when there are no prems to the subgoal *)  | 
|
282  | 
(* does a case split on the given variable *)  | 
|
283  | 
fun mk_casesplit_goal_thm ctxt (vstr,ty) gt =  | 
|
284  | 
let  | 
|
285  | 
val thy = Proof_Context.theory_of ctxt;  | 
|
286  | 
||
287  | 
val x = Free(vstr,ty);  | 
|
288  | 
val abst = Abs(vstr, ty, Term.abstract_over (x, gt));  | 
|
289  | 
||
| 60752 | 290  | 
val case_thm = case_thm_of_ty ctxt ty;  | 
| 60520 | 291  | 
|
292  | 
val abs_ct = Thm.cterm_of ctxt abst;  | 
|
293  | 
val free_ct = Thm.cterm_of ctxt x;  | 
|
294  | 
||
295  | 
val (Pv, Dv, type_insts) =  | 
|
296  | 
case (Thm.concl_of case_thm) of  | 
|
| 60521 | 297  | 
(_ $ (Pv $ (Dv as Var(_, Dty)))) =>  | 
| 60520 | 298  | 
(Pv, Dv,  | 
299  | 
Sign.typ_match thy (Dty, ty) Vartab.empty)  | 
|
300  | 
| _ => error "not a valid case thm";  | 
|
| 
60642
 
48dd1cefb4ae
simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
 
wenzelm 
parents: 
60524 
diff
changeset
 | 
301  | 
val type_cinsts = map (fn (ixn, (S, T)) => ((ixn, S), Thm.ctyp_of ctxt T))  | 
| 60520 | 302  | 
(Vartab.dest type_insts);  | 
| 
60642
 
48dd1cefb4ae
simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
 
wenzelm 
parents: 
60524 
diff
changeset
 | 
303  | 
val Pv = dest_Var (Envir.subst_term_types type_insts Pv);  | 
| 
 
48dd1cefb4ae
simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
 
wenzelm 
parents: 
60524 
diff
changeset
 | 
304  | 
val Dv = dest_Var (Envir.subst_term_types type_insts Dv);  | 
| 60520 | 305  | 
in  | 
306  | 
Conv.fconv_rule Drule.beta_eta_conversion  | 
|
307  | 
(case_thm  | 
|
| 74282 | 308  | 
|> Thm.instantiate (TVars.make type_cinsts, Vars.empty)  | 
309  | 
|> Thm.instantiate (TVars.empty, Vars.make [(Pv, abs_ct), (Dv, free_ct)]))  | 
|
| 60520 | 310  | 
end;  | 
311  | 
||
312  | 
||
313  | 
(* the find_XXX_split functions are simply doing a lightwieght (I  | 
|
314  | 
think) term matching equivalent to find where to do the next split *)  | 
|
315  | 
||
316  | 
(* assuming two twems are identical except for a free in one at a  | 
|
317  | 
subterm, or constant in another, ie assume that one term is a plit of  | 
|
318  | 
another, then gives back the free variable that has been split. *)  | 
|
319  | 
exception find_split_exp of string  | 
|
320  | 
fun find_term_split (Free v, _ $ _) = SOME v  | 
|
321  | 
| find_term_split (Free v, Const _) = SOME v  | 
|
322  | 
| find_term_split (Free v, Abs _) = SOME v (* do we really want this case? *)  | 
|
| 60521 | 323  | 
| find_term_split (Free _, Var _) = NONE (* keep searching *)  | 
| 60520 | 324  | 
| find_term_split (a $ b, a2 $ b2) =  | 
325  | 
(case find_term_split (a, a2) of  | 
|
326  | 
NONE => find_term_split (b,b2)  | 
|
327  | 
| vopt => vopt)  | 
|
| 60521 | 328  | 
| find_term_split (Abs(_,_,t1), Abs(_,_,t2)) =  | 
| 60520 | 329  | 
find_term_split (t1, t2)  | 
| 60521 | 330  | 
| find_term_split (Const (x,_), Const(x2,_)) =  | 
| 60520 | 331  | 
if x = x2 then NONE else (* keep searching *)  | 
332  | 
raise find_split_exp (* stop now *)  | 
|
333  | 
"Terms are not identical upto a free varaible! (Consts)"  | 
|
334  | 
| find_term_split (Bound i, Bound j) =  | 
|
335  | 
if i = j then NONE else (* keep searching *)  | 
|
336  | 
raise find_split_exp (* stop now *)  | 
|
337  | 
"Terms are not identical upto a free varaible! (Bound)"  | 
|
338  | 
| find_term_split _ =  | 
|
339  | 
raise find_split_exp (* stop now *)  | 
|
340  | 
"Terms are not identical upto a free varaible! (Other)";  | 
|
341  | 
||
342  | 
(* assume that "splitth" is a case split form of subgoal i of "genth",  | 
|
343  | 
then look for a free variable to split, breaking the subgoal closer to  | 
|
344  | 
splitth. *)  | 
|
345  | 
fun find_thm_split splitth i genth =  | 
|
346  | 
find_term_split (Logic.get_goal (Thm.prop_of genth) i,  | 
|
347  | 
Thm.concl_of splitth) handle find_split_exp _ => NONE;  | 
|
348  | 
||
349  | 
(* as above but searches "splitths" for a theorem that suggest a case split *)  | 
|
350  | 
fun find_thms_split splitths i genth =  | 
|
351  | 
Library.get_first (fn sth => find_thm_split sth i genth) splitths;  | 
|
352  | 
||
353  | 
||
354  | 
(* split the subgoal i of "genth" until we get to a member of  | 
|
355  | 
splitths. Assumes that genth will be a general form of splitths, that  | 
|
356  | 
can be case-split, as needed. Otherwise fails. Note: We assume that  | 
|
357  | 
all of "splitths" are split to the same level, and thus it doesn't  | 
|
358  | 
matter which one we choose to look for the next split. Simply add  | 
|
359  | 
search on splitthms and split variable, to change this. *)  | 
|
360  | 
(* Note: possible efficiency measure: when a case theorem is no longer  | 
|
361  | 
useful, drop it? *)  | 
|
362  | 
(* Note: This should not be a separate tactic but integrated into the  | 
|
363  | 
case split done during recdef's case analysis, this would avoid us  | 
|
364  | 
having to (re)search for variables to split. *)  | 
|
365  | 
fun splitto ctxt splitths genth =  | 
|
366  | 
let  | 
|
367  | 
val _ = not (null splitths) orelse error "splitto: no given splitths";  | 
|
368  | 
||
369  | 
(* check if we are a member of splitths - FIXME: quicker and  | 
|
370  | 
more flexible with discrim net. *)  | 
|
371  | 
fun solve_by_splitth th split =  | 
|
372  | 
Thm.biresolution (SOME ctxt) false [(false,split)] 1 th;  | 
|
373  | 
||
374  | 
fun split th =  | 
|
375  | 
(case find_thms_split splitths 1 th of  | 
|
376  | 
NONE =>  | 
|
377  | 
(writeln (cat_lines  | 
|
| 61268 | 378  | 
(["th:", Thm.string_of_thm ctxt th, "split ths:"] @  | 
379  | 
map (Thm.string_of_thm ctxt) splitths @ ["\n--"]));  | 
|
| 60520 | 380  | 
error "splitto: cannot find variable to split on")  | 
381  | 
| SOME v =>  | 
|
382  | 
let  | 
|
383  | 
val gt = HOLogic.dest_Trueprop (#1 (Logic.dest_implies (Thm.prop_of th)));  | 
|
384  | 
val split_thm = mk_casesplit_goal_thm ctxt v gt;  | 
|
385  | 
val (subthms, expf) = IsaND.fixed_subgoal_thms ctxt split_thm;  | 
|
386  | 
in  | 
|
387  | 
expf (map recsplitf subthms)  | 
|
388  | 
end)  | 
|
389  | 
||
390  | 
and recsplitf th =  | 
|
391  | 
(* note: multiple unifiers! we only take the first element,  | 
|
392  | 
probably fine -- there is probably only one anyway. *)  | 
|
393  | 
(case get_first (Seq.pull o solve_by_splitth th) splitths of  | 
|
394  | 
NONE => split th  | 
|
395  | 
| SOME (solved_th, _) => solved_th);  | 
|
396  | 
in  | 
|
397  | 
recsplitf genth  | 
|
398  | 
end;  | 
|
399  | 
||
400  | 
end;  | 
|
401  | 
||
402  | 
||
| 60521 | 403  | 
|
| 60520 | 404  | 
(*** basic utilities ***)  | 
405  | 
||
406  | 
structure Utils: UTILS =  | 
|
407  | 
struct  | 
|
408  | 
||
409  | 
(*standard exception for TFL*)  | 
|
410  | 
exception ERR of {module: string, func: string, mesg: string};
 | 
|
411  | 
||
412  | 
fun UTILS_ERR func mesg = ERR {module = "Utils", func = func, mesg = mesg};
 | 
|
413  | 
||
414  | 
||
| 60521 | 415  | 
fun end_itlist _ [] = raise (UTILS_ERR "end_itlist" "list too short")  | 
416  | 
| end_itlist _ [x] = x  | 
|
| 60520 | 417  | 
| end_itlist f (x :: xs) = f x (end_itlist f xs);  | 
418  | 
||
419  | 
fun itlist2 f L1 L2 base_value =  | 
|
420  | 
let fun it ([],[]) = base_value  | 
|
421  | 
| it ((a::rst1),(b::rst2)) = f a b (it (rst1,rst2))  | 
|
422  | 
| it _ = raise UTILS_ERR "itlist2" "different length lists"  | 
|
423  | 
in it (L1,L2)  | 
|
424  | 
end;  | 
|
425  | 
||
426  | 
fun pluck p =  | 
|
427  | 
let fun remv ([],_) = raise UTILS_ERR "pluck" "item not found"  | 
|
428  | 
| remv (h::t, A) = if p h then (h, rev A @ t) else remv (t,h::A)  | 
|
429  | 
in fn L => remv(L,[])  | 
|
430  | 
end;  | 
|
431  | 
||
432  | 
fun take f =  | 
|
| 60521 | 433  | 
let fun grab(0, _) = []  | 
| 60520 | 434  | 
| grab(n, x::rst) = f x::grab(n-1,rst)  | 
435  | 
in grab  | 
|
436  | 
end;  | 
|
437  | 
||
438  | 
fun zip3 [][][] = []  | 
|
439  | 
| zip3 (x::l1) (y::l2) (z::l3) = (x,y,z)::zip3 l1 l2 l3  | 
|
440  | 
| zip3 _ _ _ = raise UTILS_ERR "zip3" "different lengths";  | 
|
441  | 
||
442  | 
||
443  | 
end;  | 
|
444  | 
||
445  | 
||
| 60521 | 446  | 
|
| 60520 | 447  | 
(*** emulation of HOL's abstract syntax functions ***)  | 
448  | 
||
449  | 
structure USyntax: USYNTAX =  | 
|
450  | 
struct  | 
|
451  | 
||
452  | 
infix 4 ##;  | 
|
453  | 
||
454  | 
fun USYN_ERR func mesg = Utils.ERR {module = "USyntax", func = func, mesg = mesg};
 | 
|
455  | 
||
456  | 
||
457  | 
(*---------------------------------------------------------------------------  | 
|
458  | 
*  | 
|
459  | 
* Types  | 
|
460  | 
*  | 
|
461  | 
*---------------------------------------------------------------------------*)  | 
|
462  | 
val mk_prim_vartype = TVar;  | 
|
| 69593 | 463  | 
fun mk_vartype s = mk_prim_vartype ((s, 0), \<^sort>\<open>type\<close>);  | 
| 60520 | 464  | 
|
465  | 
(* But internally, it's useful *)  | 
|
466  | 
fun dest_vtype (TVar x) = x  | 
|
467  | 
| dest_vtype _ = raise USYN_ERR "dest_vtype" "not a flexible type variable";  | 
|
468  | 
||
469  | 
val is_vartype = can dest_vtype;  | 
|
470  | 
||
471  | 
val type_vars = map mk_prim_vartype o Misc_Legacy.typ_tvars  | 
|
472  | 
fun type_varsl L = distinct (op =) (fold (curry op @ o type_vars) L []);  | 
|
473  | 
||
474  | 
val alpha = mk_vartype "'a"  | 
|
475  | 
||
476  | 
val strip_prod_type = HOLogic.flatten_tupleT;  | 
|
477  | 
||
478  | 
||
479  | 
||
480  | 
(*---------------------------------------------------------------------------  | 
|
481  | 
*  | 
|
482  | 
* Terms  | 
|
483  | 
*  | 
|
484  | 
*---------------------------------------------------------------------------*)  | 
|
485  | 
||
486  | 
(* Free variables, in order of occurrence, from left to right in the  | 
|
487  | 
* syntax tree. *)  | 
|
488  | 
fun free_vars_lr tm =  | 
|
489  | 
let fun memb x = let fun m[] = false | m(y::rst) = (x=y)orelse m rst in m end  | 
|
490  | 
fun add (t, frees) = case t of  | 
|
491  | 
Free _ => if (memb t frees) then frees else t::frees  | 
|
492  | 
| Abs (_,_,body) => add(body,frees)  | 
|
493  | 
| f$t => add(t, add(f, frees))  | 
|
494  | 
| _ => frees  | 
|
495  | 
in rev(add(tm,[]))  | 
|
496  | 
end;  | 
|
497  | 
||
498  | 
||
499  | 
||
500  | 
val type_vars_in_term = map mk_prim_vartype o Misc_Legacy.term_tvars;  | 
|
501  | 
||
502  | 
||
503  | 
||
504  | 
(* Prelogic *)  | 
|
505  | 
fun dest_tybinding (v,ty) = (#1(dest_vtype v),ty)  | 
|
506  | 
fun inst theta = subst_vars (map dest_tybinding theta,[])  | 
|
507  | 
||
508  | 
||
509  | 
(* Construction routines *)  | 
|
510  | 
||
511  | 
fun mk_abs{Bvar as Var((s,_),ty),Body}  = Abs(s,ty,abstract_over(Bvar,Body))
 | 
|
512  | 
  | mk_abs{Bvar as Free(s,ty),Body}  = Abs(s,ty,abstract_over(Bvar,Body))
 | 
|
513  | 
| mk_abs _ = raise USYN_ERR "mk_abs" "Bvar is not a variable";  | 
|
514  | 
||
515  | 
||
516  | 
fun mk_imp{ant,conseq} =
 | 
|
| 69593 | 517  | 
let val c = Const(\<^const_name>\<open>HOL.implies\<close>,HOLogic.boolT --> HOLogic.boolT --> HOLogic.boolT)  | 
| 60520 | 518  | 
in list_comb(c,[ant,conseq])  | 
519  | 
end;  | 
|
520  | 
||
521  | 
fun mk_select (r as {Bvar,Body}) =
 | 
|
522  | 
let val ty = type_of Bvar  | 
|
| 69593 | 523  | 
val c = Const(\<^const_name>\<open>Eps\<close>,(ty --> HOLogic.boolT) --> ty)  | 
| 60520 | 524  | 
in list_comb(c,[mk_abs r])  | 
525  | 
end;  | 
|
526  | 
||
527  | 
fun mk_forall (r as {Bvar,Body}) =
 | 
|
528  | 
let val ty = type_of Bvar  | 
|
| 69593 | 529  | 
val c = Const(\<^const_name>\<open>All\<close>,(ty --> HOLogic.boolT) --> HOLogic.boolT)  | 
| 60520 | 530  | 
in list_comb(c,[mk_abs r])  | 
531  | 
end;  | 
|
532  | 
||
533  | 
fun mk_exists (r as {Bvar,Body}) =
 | 
|
534  | 
let val ty = type_of Bvar  | 
|
| 69593 | 535  | 
val c = Const(\<^const_name>\<open>Ex\<close>,(ty --> HOLogic.boolT) --> HOLogic.boolT)  | 
| 60520 | 536  | 
in list_comb(c,[mk_abs r])  | 
537  | 
end;  | 
|
538  | 
||
539  | 
||
540  | 
fun mk_conj{conj1,conj2} =
 | 
|
| 69593 | 541  | 
let val c = Const(\<^const_name>\<open>HOL.conj\<close>,HOLogic.boolT --> HOLogic.boolT --> HOLogic.boolT)  | 
| 60520 | 542  | 
in list_comb(c,[conj1,conj2])  | 
543  | 
end;  | 
|
544  | 
||
545  | 
fun mk_disj{disj1,disj2} =
 | 
|
| 69593 | 546  | 
let val c = Const(\<^const_name>\<open>HOL.disj\<close>,HOLogic.boolT --> HOLogic.boolT --> HOLogic.boolT)  | 
| 60520 | 547  | 
in list_comb(c,[disj1,disj2])  | 
548  | 
end;  | 
|
549  | 
||
550  | 
fun prod_ty ty1 ty2 = HOLogic.mk_prodT (ty1,ty2);  | 
|
551  | 
||
552  | 
local  | 
|
553  | 
fun mk_uncurry (xt, yt, zt) =  | 
|
| 69593 | 554  | 
Const(\<^const_name>\<open>case_prod\<close>, (xt --> yt --> zt) --> prod_ty xt yt --> zt)  | 
555  | 
fun dest_pair(Const(\<^const_name>\<open>Pair\<close>,_) $ M $ N) = {fst=M, snd=N}
 | 
|
| 60520 | 556  | 
| dest_pair _ = raise USYN_ERR "dest_pair" "not a pair"  | 
557  | 
fun is_var (Var _) = true | is_var (Free _) = true | is_var _ = false  | 
|
558  | 
in  | 
|
559  | 
fun mk_pabs{varstruct,body} =
 | 
|
560  | 
let fun mpa (varstruct, body) =  | 
|
561  | 
if is_var varstruct  | 
|
562  | 
       then mk_abs {Bvar = varstruct, Body = body}
 | 
|
563  | 
       else let val {fst, snd} = dest_pair varstruct
 | 
|
564  | 
in mk_uncurry (type_of fst, type_of snd, type_of body) $  | 
|
565  | 
mpa (fst, mpa (snd, body))  | 
|
566  | 
end  | 
|
567  | 
in mpa (varstruct, body) end  | 
|
568  | 
handle TYPE _ => raise USYN_ERR "mk_pabs" "";  | 
|
569  | 
end;  | 
|
570  | 
||
571  | 
(* Destruction routines *)  | 
|
572  | 
||
573  | 
datatype lambda = VAR   of {Name : string, Ty : typ}
 | 
|
574  | 
                | CONST of {Name : string, Ty : typ}
 | 
|
575  | 
                | COMB  of {Rator: term, Rand : term}
 | 
|
576  | 
                | LAMB  of {Bvar : term, Body : term};
 | 
|
577  | 
||
578  | 
||
| 60521 | 579  | 
fun dest_term(Var((s,_),ty)) = VAR{Name = s, Ty = ty}
 | 
| 60520 | 580  | 
  | dest_term(Free(s,ty))    = VAR{Name = s, Ty = ty}
 | 
581  | 
  | dest_term(Const(s,ty))   = CONST{Name = s, Ty = ty}
 | 
|
582  | 
  | dest_term(M$N)           = COMB{Rator=M,Rand=N}
 | 
|
583  | 
| dest_term(Abs(s,ty,M)) = let val v = Free(s,ty)  | 
|
584  | 
                               in LAMB{Bvar = v, Body = Term.betapply (M,v)}
 | 
|
585  | 
end  | 
|
586  | 
| dest_term(Bound _) = raise USYN_ERR "dest_term" "Bound";  | 
|
587  | 
||
588  | 
fun dest_const(Const(s,ty)) = {Name = s, Ty = ty}
 | 
|
589  | 
| dest_const _ = raise USYN_ERR "dest_const" "not a constant";  | 
|
590  | 
||
591  | 
fun dest_comb(t1 $ t2) = {Rator = t1, Rand = t2}
 | 
|
592  | 
| dest_comb _ = raise USYN_ERR "dest_comb" "not a comb";  | 
|
593  | 
||
| 60524 | 594  | 
fun dest_abs used (a as Abs(s, ty, _)) =  | 
| 60520 | 595  | 
let  | 
596  | 
val s' = singleton (Name.variant_list used) s;  | 
|
597  | 
val v = Free(s', ty);  | 
|
598  | 
     in ({Bvar = v, Body = Term.betapply (a,v)}, s'::used)
 | 
|
599  | 
end  | 
|
600  | 
| dest_abs _ _ = raise USYN_ERR "dest_abs" "not an abstraction";  | 
|
601  | 
||
| 69593 | 602  | 
fun dest_eq(Const(\<^const_name>\<open>HOL.eq\<close>,_) $ M $ N) = {lhs=M, rhs=N}
 | 
| 60520 | 603  | 
| dest_eq _ = raise USYN_ERR "dest_eq" "not an equality";  | 
604  | 
||
| 69593 | 605  | 
fun dest_imp(Const(\<^const_name>\<open>HOL.implies\<close>,_) $ M $ N) = {ant=M, conseq=N}
 | 
| 60520 | 606  | 
| dest_imp _ = raise USYN_ERR "dest_imp" "not an implication";  | 
607  | 
||
| 69593 | 608  | 
fun dest_forall(Const(\<^const_name>\<open>All\<close>,_) $ (a as Abs _)) = fst (dest_abs [] a)  | 
| 60520 | 609  | 
| dest_forall _ = raise USYN_ERR "dest_forall" "not a forall";  | 
610  | 
||
| 69593 | 611  | 
fun dest_exists(Const(\<^const_name>\<open>Ex\<close>,_) $ (a as Abs _)) = fst (dest_abs [] a)  | 
| 60520 | 612  | 
| dest_exists _ = raise USYN_ERR "dest_exists" "not an existential";  | 
613  | 
||
| 69593 | 614  | 
fun dest_neg(Const(\<^const_name>\<open>Not\<close>,_) $ M) = M  | 
| 60520 | 615  | 
| dest_neg _ = raise USYN_ERR "dest_neg" "not a negation";  | 
616  | 
||
| 69593 | 617  | 
fun dest_conj(Const(\<^const_name>\<open>HOL.conj\<close>,_) $ M $ N) = {conj1=M, conj2=N}
 | 
| 60520 | 618  | 
| dest_conj _ = raise USYN_ERR "dest_conj" "not a conjunction";  | 
619  | 
||
| 69593 | 620  | 
fun dest_disj(Const(\<^const_name>\<open>HOL.disj\<close>,_) $ M $ N) = {disj1=M, disj2=N}
 | 
| 60520 | 621  | 
| dest_disj _ = raise USYN_ERR "dest_disj" "not a disjunction";  | 
622  | 
||
623  | 
fun mk_pair{fst,snd} =
 | 
|
624  | 
let val ty1 = type_of fst  | 
|
625  | 
val ty2 = type_of snd  | 
|
| 69593 | 626  | 
val c = Const(\<^const_name>\<open>Pair\<close>,ty1 --> ty2 --> prod_ty ty1 ty2)  | 
| 60520 | 627  | 
in list_comb(c,[fst,snd])  | 
628  | 
end;  | 
|
629  | 
||
| 69593 | 630  | 
fun dest_pair(Const(\<^const_name>\<open>Pair\<close>,_) $ M $ N) = {fst=M, snd=N}
 | 
| 60520 | 631  | 
| dest_pair _ = raise USYN_ERR "dest_pair" "not a pair";  | 
632  | 
||
633  | 
||
| 69593 | 634  | 
local fun ucheck t = (if #Name (dest_const t) = \<^const_name>\<open>case_prod\<close> then t  | 
| 60520 | 635  | 
else raise Match)  | 
636  | 
in  | 
|
637  | 
fun dest_pabs used tm =  | 
|
638  | 
   let val ({Bvar,Body}, used') = dest_abs used tm
 | 
|
639  | 
   in {varstruct = Bvar, body = Body, used = used'}
 | 
|
640  | 
end handle Utils.ERR _ =>  | 
|
641  | 
          let val {Rator,Rand} = dest_comb tm
 | 
|
642  | 
val _ = ucheck Rator  | 
|
643  | 
              val {varstruct = lv, body, used = used'} = dest_pabs used Rand
 | 
|
644  | 
              val {varstruct = rv, body, used = used''} = dest_pabs used' body
 | 
|
645  | 
          in {varstruct = mk_pair {fst = lv, snd = rv}, body = body, used = used''}
 | 
|
646  | 
end  | 
|
647  | 
end;  | 
|
648  | 
||
649  | 
||
650  | 
val lhs = #lhs o dest_eq  | 
|
651  | 
val rhs = #rhs o dest_eq  | 
|
652  | 
val rand = #Rand o dest_comb  | 
|
653  | 
||
654  | 
||
655  | 
(* Query routines *)  | 
|
656  | 
val is_imp = can dest_imp  | 
|
657  | 
val is_forall = can dest_forall  | 
|
658  | 
val is_exists = can dest_exists  | 
|
659  | 
val is_neg = can dest_neg  | 
|
660  | 
val is_conj = can dest_conj  | 
|
661  | 
val is_disj = can dest_disj  | 
|
662  | 
val is_pair = can dest_pair  | 
|
663  | 
val is_pabs = can (dest_pabs [])  | 
|
664  | 
||
665  | 
||
666  | 
(* Construction of a cterm from a list of Terms *)  | 
|
667  | 
||
668  | 
fun list_mk_abs(L,tm) = fold_rev (fn v => fn M => mk_abs{Bvar=v, Body=M}) L tm;
 | 
|
669  | 
||
670  | 
(* These others are almost never used *)  | 
|
671  | 
fun list_mk_imp(A,c) = fold_rev (fn a => fn tm => mk_imp{ant=a,conseq=tm}) A c;
 | 
|
672  | 
fun list_mk_forall(V,t) = fold_rev (fn v => fn b => mk_forall{Bvar=v, Body=b})V t;
 | 
|
673  | 
val list_mk_conj = Utils.end_itlist(fn c1 => fn tm => mk_conj{conj1=c1, conj2=tm})
 | 
|
674  | 
||
675  | 
||
676  | 
(* Need to reverse? *)  | 
|
677  | 
fun gen_all tm = list_mk_forall(Misc_Legacy.term_frees tm, tm);  | 
|
678  | 
||
679  | 
(* Destructing a cterm to a list of Terms *)  | 
|
680  | 
fun strip_comb tm =  | 
|
681  | 
let fun dest(M$N, A) = dest(M, N::A)  | 
|
682  | 
| dest x = x  | 
|
683  | 
in dest(tm,[])  | 
|
684  | 
end;  | 
|
685  | 
||
686  | 
fun strip_abs(tm as Abs _) =  | 
|
687  | 
       let val ({Bvar,Body}, _) = dest_abs [] tm
 | 
|
688  | 
val (bvs, core) = strip_abs Body  | 
|
689  | 
in (Bvar::bvs, core)  | 
|
690  | 
end  | 
|
691  | 
| strip_abs M = ([],M);  | 
|
692  | 
||
693  | 
||
694  | 
fun strip_imp fm =  | 
|
695  | 
if (is_imp fm)  | 
|
696  | 
   then let val {ant,conseq} = dest_imp fm
 | 
|
697  | 
val (was,wb) = strip_imp conseq  | 
|
698  | 
in ((ant::was), wb)  | 
|
699  | 
end  | 
|
700  | 
else ([],fm);  | 
|
701  | 
||
702  | 
fun strip_forall fm =  | 
|
703  | 
if (is_forall fm)  | 
|
704  | 
   then let val {Bvar,Body} = dest_forall fm
 | 
|
705  | 
val (bvs,core) = strip_forall Body  | 
|
706  | 
in ((Bvar::bvs), core)  | 
|
707  | 
end  | 
|
708  | 
else ([],fm);  | 
|
709  | 
||
710  | 
||
711  | 
fun strip_exists fm =  | 
|
712  | 
if (is_exists fm)  | 
|
713  | 
   then let val {Bvar, Body} = dest_exists fm
 | 
|
714  | 
val (bvs,core) = strip_exists Body  | 
|
715  | 
in (Bvar::bvs, core)  | 
|
716  | 
end  | 
|
717  | 
else ([],fm);  | 
|
718  | 
||
719  | 
fun strip_disj w =  | 
|
720  | 
if (is_disj w)  | 
|
721  | 
   then let val {disj1,disj2} = dest_disj w
 | 
|
722  | 
in (strip_disj disj1@strip_disj disj2)  | 
|
723  | 
end  | 
|
724  | 
else [w];  | 
|
725  | 
||
726  | 
||
727  | 
(* Miscellaneous *)  | 
|
728  | 
||
729  | 
fun mk_vstruct ty V =  | 
|
| 69593 | 730  | 
let fun follow_prod_type (Type(\<^type_name>\<open>Product_Type.prod\<close>,[ty1,ty2])) vs =  | 
| 60520 | 731  | 
let val (ltm,vs1) = follow_prod_type ty1 vs  | 
732  | 
val (rtm,vs2) = follow_prod_type ty2 vs1  | 
|
733  | 
              in (mk_pair{fst=ltm, snd=rtm}, vs2) end
 | 
|
734  | 
| follow_prod_type _ (v::vs) = (v,vs)  | 
|
735  | 
in #1 (follow_prod_type ty V) end;  | 
|
736  | 
||
737  | 
||
738  | 
(* Search a term for a sub-term satisfying the predicate p. *)  | 
|
739  | 
fun find_term p =  | 
|
740  | 
let fun find tm =  | 
|
741  | 
if (p tm) then SOME tm  | 
|
742  | 
else case tm of  | 
|
743  | 
Abs(_,_,body) => find body  | 
|
744  | 
| (t$u) => (case find t of NONE => find u | some => some)  | 
|
745  | 
| _ => NONE  | 
|
746  | 
in find  | 
|
747  | 
end;  | 
|
748  | 
||
749  | 
fun dest_relation tm =  | 
|
750  | 
if (type_of tm = HOLogic.boolT)  | 
|
| 69593 | 751  | 
then let val (Const(\<^const_name>\<open>Set.member\<close>,_) $ (Const(\<^const_name>\<open>Pair\<close>,_)$y$x) $ R) = tm  | 
| 60520 | 752  | 
in (R,y,x)  | 
753  | 
end handle Bind => raise USYN_ERR "dest_relation" "unexpected term structure"  | 
|
754  | 
else raise USYN_ERR "dest_relation" "not a boolean term";  | 
|
755  | 
||
| 69593 | 756  | 
fun is_WFR (Const(\<^const_name>\<open>Wellfounded.wf\<close>,_)$_) = true  | 
| 60520 | 757  | 
| is_WFR _ = false;  | 
758  | 
||
759  | 
fun ARB ty = mk_select{Bvar=Free("v",ty),
 | 
|
| 69593 | 760  | 
Body=Const(\<^const_name>\<open>True\<close>,HOLogic.boolT)};  | 
| 60520 | 761  | 
|
762  | 
end;  | 
|
763  | 
||
764  | 
||
| 60521 | 765  | 
|
| 60520 | 766  | 
(*** derived cterm destructors ***)  | 
767  | 
||
768  | 
structure Dcterm: DCTERM =  | 
|
769  | 
struct  | 
|
770  | 
||
771  | 
fun ERR func mesg = Utils.ERR {module = "Dcterm", func = func, mesg = mesg};
 | 
|
772  | 
||
773  | 
||
774  | 
fun dest_comb t = Thm.dest_comb t  | 
|
775  | 
handle CTERM (msg, _) => raise ERR "dest_comb" msg;  | 
|
776  | 
||
| 
74525
 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 
wenzelm 
parents: 
74518 
diff
changeset
 | 
777  | 
fun dest_abs t = Thm.dest_abs_global t  | 
| 60520 | 778  | 
handle CTERM (msg, _) => raise ERR "dest_abs" msg;  | 
779  | 
||
780  | 
fun capply t u = Thm.apply t u  | 
|
781  | 
handle CTERM (msg, _) => raise ERR "capply" msg;  | 
|
782  | 
||
783  | 
fun cabs a t = Thm.lambda a t  | 
|
784  | 
handle CTERM (msg, _) => raise ERR "cabs" msg;  | 
|
785  | 
||
786  | 
||
787  | 
(*---------------------------------------------------------------------------  | 
|
788  | 
* Some simple constructor functions.  | 
|
789  | 
*---------------------------------------------------------------------------*)  | 
|
790  | 
||
| 69593 | 791  | 
val mk_hol_const = Thm.cterm_of \<^theory_context>\<open>HOL\<close> o Const;  | 
| 60520 | 792  | 
|
793  | 
fun mk_exists (r as (Bvar, Body)) =  | 
|
794  | 
let val ty = Thm.typ_of_cterm Bvar  | 
|
| 69593 | 795  | 
val c = mk_hol_const(\<^const_name>\<open>Ex\<close>, (ty --> HOLogic.boolT) --> HOLogic.boolT)  | 
| 60520 | 796  | 
in capply c (uncurry cabs r) end;  | 
797  | 
||
798  | 
||
| 69593 | 799  | 
local val c = mk_hol_const(\<^const_name>\<open>HOL.conj\<close>, HOLogic.boolT --> HOLogic.boolT --> HOLogic.boolT)  | 
| 60520 | 800  | 
in fun mk_conj(conj1,conj2) = capply (capply c conj1) conj2  | 
801  | 
end;  | 
|
802  | 
||
| 69593 | 803  | 
local val c = mk_hol_const(\<^const_name>\<open>HOL.disj\<close>, HOLogic.boolT --> HOLogic.boolT --> HOLogic.boolT)  | 
| 60520 | 804  | 
in fun mk_disj(disj1,disj2) = capply (capply c disj1) disj2  | 
805  | 
end;  | 
|
806  | 
||
807  | 
||
808  | 
(*---------------------------------------------------------------------------  | 
|
809  | 
* The primitives.  | 
|
810  | 
*---------------------------------------------------------------------------*)  | 
|
811  | 
fun dest_const ctm =  | 
|
812  | 
(case Thm.term_of ctm  | 
|
813  | 
      of Const(s,ty) => {Name = s, Ty = ty}
 | 
|
814  | 
| _ => raise ERR "dest_const" "not a constant");  | 
|
815  | 
||
816  | 
fun dest_var ctm =  | 
|
817  | 
(case Thm.term_of ctm  | 
|
| 60521 | 818  | 
      of Var((s,_),ty) => {Name=s, Ty=ty}
 | 
| 60520 | 819  | 
       | Free(s,ty)    => {Name=s, Ty=ty}
 | 
820  | 
| _ => raise ERR "dest_var" "not a variable");  | 
|
821  | 
||
822  | 
||
823  | 
(*---------------------------------------------------------------------------  | 
|
824  | 
* Derived destructor operations.  | 
|
825  | 
*---------------------------------------------------------------------------*)  | 
|
826  | 
||
827  | 
fun dest_monop expected tm =  | 
|
828  | 
let  | 
|
829  | 
   fun err () = raise ERR "dest_monop" ("Not a(n) " ^ quote expected);
 | 
|
830  | 
val (c, N) = dest_comb tm handle Utils.ERR _ => err ();  | 
|
831  | 
val name = #Name (dest_const c handle Utils.ERR _ => err ());  | 
|
832  | 
in if name = expected then N else err () end;  | 
|
833  | 
||
834  | 
fun dest_binop expected tm =  | 
|
835  | 
let  | 
|
836  | 
   fun err () = raise ERR "dest_binop" ("Not a(n) " ^ quote expected);
 | 
|
837  | 
val (M, N) = dest_comb tm handle Utils.ERR _ => err ()  | 
|
838  | 
in (dest_monop expected M, N) handle Utils.ERR _ => err () end;  | 
|
839  | 
||
840  | 
fun dest_binder expected tm =  | 
|
| 74518 | 841  | 
dest_abs (dest_monop expected tm)  | 
| 60520 | 842  | 
  handle Utils.ERR _ => raise ERR "dest_binder" ("Not a(n) " ^ quote expected);
 | 
843  | 
||
844  | 
||
| 69593 | 845  | 
val dest_neg = dest_monop \<^const_name>\<open>Not\<close>  | 
846  | 
val dest_pair = dest_binop \<^const_name>\<open>Pair\<close>  | 
|
847  | 
val dest_eq = dest_binop \<^const_name>\<open>HOL.eq\<close>  | 
|
848  | 
val dest_imp = dest_binop \<^const_name>\<open>HOL.implies\<close>  | 
|
849  | 
val dest_conj = dest_binop \<^const_name>\<open>HOL.conj\<close>  | 
|
850  | 
val dest_disj = dest_binop \<^const_name>\<open>HOL.disj\<close>  | 
|
851  | 
val dest_exists = dest_binder \<^const_name>\<open>Ex\<close>  | 
|
852  | 
val dest_forall = dest_binder \<^const_name>\<open>All\<close>  | 
|
| 60520 | 853  | 
|
854  | 
(* Query routines *)  | 
|
855  | 
||
856  | 
val is_eq = can dest_eq  | 
|
857  | 
val is_imp = can dest_imp  | 
|
858  | 
val is_forall = can dest_forall  | 
|
859  | 
val is_exists = can dest_exists  | 
|
860  | 
val is_neg = can dest_neg  | 
|
861  | 
val is_conj = can dest_conj  | 
|
862  | 
val is_disj = can dest_disj  | 
|
863  | 
val is_pair = can dest_pair  | 
|
864  | 
||
865  | 
||
866  | 
(*---------------------------------------------------------------------------  | 
|
867  | 
* Iterated creation.  | 
|
868  | 
*---------------------------------------------------------------------------*)  | 
|
869  | 
val list_mk_disj = Utils.end_itlist (fn d1 => fn tm => mk_disj (d1, tm));  | 
|
870  | 
||
871  | 
(*---------------------------------------------------------------------------  | 
|
872  | 
* Iterated destruction. (To the "right" in a term.)  | 
|
873  | 
*---------------------------------------------------------------------------*)  | 
|
874  | 
fun strip break tm =  | 
|
875  | 
let fun dest (p as (ctm,accum)) =  | 
|
876  | 
let val (M,N) = break ctm  | 
|
877  | 
in dest (N, M::accum)  | 
|
878  | 
end handle Utils.ERR _ => p  | 
|
879  | 
in dest (tm,[])  | 
|
880  | 
end;  | 
|
881  | 
||
882  | 
fun rev2swap (x,l) = (rev l, x);  | 
|
883  | 
||
884  | 
val strip_comb = strip (Library.swap o dest_comb) (* Goes to the "left" *)  | 
|
885  | 
val strip_imp = rev2swap o strip dest_imp  | 
|
| 74518 | 886  | 
val strip_abs = rev2swap o strip dest_abs  | 
| 60520 | 887  | 
val strip_forall = rev2swap o strip dest_forall  | 
888  | 
val strip_exists = rev2swap o strip dest_exists  | 
|
889  | 
||
890  | 
val strip_disj = rev o (op::) o strip dest_disj  | 
|
891  | 
||
892  | 
||
893  | 
(*---------------------------------------------------------------------------  | 
|
894  | 
* Going into and out of prop  | 
|
895  | 
*---------------------------------------------------------------------------*)  | 
|
896  | 
||
897  | 
fun is_Trueprop ct =  | 
|
898  | 
(case Thm.term_of ct of  | 
|
| 69593 | 899  | 
Const (\<^const_name>\<open>Trueprop\<close>, _) $ _ => true  | 
| 60520 | 900  | 
| _ => false);  | 
901  | 
||
| 69593 | 902  | 
fun mk_prop ct = if is_Trueprop ct then ct else Thm.apply \<^cterm>\<open>Trueprop\<close> ct;  | 
| 60520 | 903  | 
fun drop_prop ct = if is_Trueprop ct then Thm.dest_arg ct else ct;  | 
904  | 
||
905  | 
end;  | 
|
906  | 
||
907  | 
||
| 60521 | 908  | 
|
| 60520 | 909  | 
(*** emulation of HOL inference rules for TFL ***)  | 
910  | 
||
911  | 
structure Rules: RULES =  | 
|
912  | 
struct  | 
|
913  | 
||
914  | 
fun RULES_ERR func mesg = Utils.ERR {module = "Rules", func = func, mesg = mesg};
 | 
|
915  | 
||
916  | 
||
| 60949 | 917  | 
fun cconcl thm = Dcterm.drop_prop (Thm.cprop_of thm);  | 
918  | 
fun chyps thm = map Dcterm.drop_prop (Thm.chyps_of thm);  | 
|
| 60520 | 919  | 
|
920  | 
fun dest_thm thm =  | 
|
| 61038 | 921  | 
(map HOLogic.dest_Trueprop (Thm.hyps_of thm), HOLogic.dest_Trueprop (Thm.prop_of thm))  | 
922  | 
handle TERM _ => raise RULES_ERR "dest_thm" "missing Trueprop";  | 
|
| 60520 | 923  | 
|
924  | 
||
925  | 
(* Inference rules *)  | 
|
926  | 
||
927  | 
(*---------------------------------------------------------------------------  | 
|
928  | 
* Equality (one step)  | 
|
929  | 
*---------------------------------------------------------------------------*)  | 
|
930  | 
||
| 
67710
 
cc2db3239932
added HOLogic.mk_obj_eq convenience and eliminated some clones;
 
wenzelm 
parents: 
66251 
diff
changeset
 | 
931  | 
fun REFL tm = HOLogic.mk_obj_eq (Thm.reflexive tm)  | 
| 60520 | 932  | 
handle THM (msg, _, _) => raise RULES_ERR "REFL" msg;  | 
933  | 
||
934  | 
fun SYM thm = thm RS sym  | 
|
935  | 
handle THM (msg, _, _) => raise RULES_ERR "SYM" msg;  | 
|
936  | 
||
937  | 
fun ALPHA thm ctm1 =  | 
|
938  | 
let  | 
|
939  | 
val ctm2 = Thm.cprop_of thm;  | 
|
940  | 
val ctm2_eq = Thm.reflexive ctm2;  | 
|
941  | 
val ctm1_eq = Thm.reflexive ctm1;  | 
|
942  | 
in Thm.equal_elim (Thm.transitive ctm2_eq ctm1_eq) thm end  | 
|
943  | 
handle THM (msg, _, _) => raise RULES_ERR "ALPHA" msg;  | 
|
944  | 
||
945  | 
fun rbeta th =  | 
|
946  | 
(case Dcterm.strip_comb (cconcl th) of  | 
|
| 60521 | 947  | 
(_, [_, r]) => Thm.transitive th (Thm.beta_conversion false r)  | 
| 60520 | 948  | 
| _ => raise RULES_ERR "rbeta" "");  | 
949  | 
||
950  | 
||
951  | 
(*----------------------------------------------------------------------------  | 
|
952  | 
* Implication and the assumption list  | 
|
953  | 
*  | 
|
954  | 
* Assumptions get stuck on the meta-language assumption list. Implications  | 
|
955  | 
* are in the object language, so discharging an assumption "A" from theorem  | 
|
956  | 
* "B" results in something that looks like "A --> B".  | 
|
957  | 
*---------------------------------------------------------------------------*)  | 
|
958  | 
||
959  | 
fun ASSUME ctm = Thm.assume (Dcterm.mk_prop ctm);  | 
|
960  | 
||
961  | 
||
962  | 
(*---------------------------------------------------------------------------  | 
|
963  | 
* Implication in TFL is -->. Meta-language implication (==>) is only used  | 
|
964  | 
* in the implementation of some of the inference rules below.  | 
|
965  | 
*---------------------------------------------------------------------------*)  | 
|
966  | 
fun MP th1 th2 = th2 RS (th1 RS mp)  | 
|
967  | 
handle THM (msg, _, _) => raise RULES_ERR "MP" msg;  | 
|
968  | 
||
969  | 
(*forces the first argument to be a proposition if necessary*)  | 
|
970  | 
fun DISCH tm thm = Thm.implies_intr (Dcterm.mk_prop tm) thm COMP impI  | 
|
971  | 
handle THM (msg, _, _) => raise RULES_ERR "DISCH" msg;  | 
|
972  | 
||
| 60949 | 973  | 
fun DISCH_ALL thm = fold_rev DISCH (Thm.chyps_of thm) thm;  | 
| 60520 | 974  | 
|
975  | 
||
976  | 
fun FILTER_DISCH_ALL P thm =  | 
|
977  | 
let fun check tm = P (Thm.term_of tm)  | 
|
978  | 
in fold_rev (fn tm => fn th => if check tm then DISCH tm th else th) (chyps thm) thm  | 
|
979  | 
end;  | 
|
980  | 
||
981  | 
fun UNDISCH thm =  | 
|
982  | 
let val tm = Dcterm.mk_prop (#1 (Dcterm.dest_imp (cconcl thm)))  | 
|
983  | 
in Thm.implies_elim (thm RS mp) (ASSUME tm) end  | 
|
984  | 
handle Utils.ERR _ => raise RULES_ERR "UNDISCH" ""  | 
|
985  | 
| THM _ => raise RULES_ERR "UNDISCH" "";  | 
|
986  | 
||
987  | 
fun PROVE_HYP ath bth = MP (DISCH (cconcl ath) bth) ath;  | 
|
988  | 
||
| 60522 | 989  | 
fun IMP_TRANS th1 th2 = th2 RS (th1 RS @{thm tfl_imp_trans})
 | 
| 60520 | 990  | 
handle THM (msg, _, _) => raise RULES_ERR "IMP_TRANS" msg;  | 
991  | 
||
992  | 
||
993  | 
(*----------------------------------------------------------------------------  | 
|
994  | 
* Conjunction  | 
|
995  | 
*---------------------------------------------------------------------------*)  | 
|
996  | 
||
997  | 
fun CONJUNCT1 thm = thm RS conjunct1  | 
|
998  | 
handle THM (msg, _, _) => raise RULES_ERR "CONJUNCT1" msg;  | 
|
999  | 
||
1000  | 
fun CONJUNCT2 thm = thm RS conjunct2  | 
|
1001  | 
handle THM (msg, _, _) => raise RULES_ERR "CONJUNCT2" msg;  | 
|
1002  | 
||
1003  | 
fun CONJUNCTS th = CONJUNCTS (CONJUNCT1 th) @ CONJUNCTS (CONJUNCT2 th) handle Utils.ERR _ => [th];  | 
|
1004  | 
||
1005  | 
fun LIST_CONJ [] = raise RULES_ERR "LIST_CONJ" "empty list"  | 
|
1006  | 
| LIST_CONJ [th] = th  | 
|
1007  | 
| LIST_CONJ (th :: rst) = MP (MP (conjI COMP (impI RS impI)) th) (LIST_CONJ rst)  | 
|
1008  | 
handle THM (msg, _, _) => raise RULES_ERR "LIST_CONJ" msg;  | 
|
1009  | 
||
1010  | 
||
1011  | 
(*----------------------------------------------------------------------------  | 
|
1012  | 
* Disjunction  | 
|
1013  | 
*---------------------------------------------------------------------------*)  | 
|
1014  | 
local  | 
|
1015  | 
val prop = Thm.prop_of disjI1  | 
|
| 60524 | 1016  | 
val [_,Q] = Misc_Legacy.term_vars prop  | 
| 69593 | 1017  | 
val disj1 = Thm.forall_intr (Thm.cterm_of \<^context> Q) disjI1  | 
| 60520 | 1018  | 
in  | 
1019  | 
fun DISJ1 thm tm = thm RS (Thm.forall_elim (Dcterm.drop_prop tm) disj1)  | 
|
1020  | 
handle THM (msg, _, _) => raise RULES_ERR "DISJ1" msg;  | 
|
1021  | 
end;  | 
|
1022  | 
||
1023  | 
local  | 
|
1024  | 
val prop = Thm.prop_of disjI2  | 
|
| 60521 | 1025  | 
val [P,_] = Misc_Legacy.term_vars prop  | 
| 69593 | 1026  | 
val disj2 = Thm.forall_intr (Thm.cterm_of \<^context> P) disjI2  | 
| 60520 | 1027  | 
in  | 
1028  | 
fun DISJ2 tm thm = thm RS (Thm.forall_elim (Dcterm.drop_prop tm) disj2)  | 
|
1029  | 
handle THM (msg, _, _) => raise RULES_ERR "DISJ2" msg;  | 
|
1030  | 
end;  | 
|
1031  | 
||
1032  | 
||
1033  | 
(*----------------------------------------------------------------------------  | 
|
1034  | 
*  | 
|
1035  | 
* A1 |- M1, ..., An |- Mn  | 
|
1036  | 
* ---------------------------------------------------  | 
|
1037  | 
* [A1 |- M1 \/ ... \/ Mn, ..., An |- M1 \/ ... \/ Mn]  | 
|
1038  | 
*  | 
|
1039  | 
*---------------------------------------------------------------------------*)  | 
|
1040  | 
||
1041  | 
||
1042  | 
fun EVEN_ORS thms =  | 
|
1043  | 
let fun blue ldisjs [] _ = []  | 
|
1044  | 
| blue ldisjs (th::rst) rdisjs =  | 
|
1045  | 
let val tail = tl rdisjs  | 
|
1046  | 
val rdisj_tl = Dcterm.list_mk_disj tail  | 
|
1047  | 
in fold_rev DISJ2 ldisjs (DISJ1 th rdisj_tl)  | 
|
1048  | 
:: blue (ldisjs @ [cconcl th]) rst tail  | 
|
1049  | 
end handle Utils.ERR _ => [fold_rev DISJ2 ldisjs th]  | 
|
1050  | 
in blue [] thms (map cconcl thms) end;  | 
|
1051  | 
||
1052  | 
||
1053  | 
(*----------------------------------------------------------------------------  | 
|
1054  | 
*  | 
|
1055  | 
* A |- P \/ Q B,P |- R C,Q |- R  | 
|
1056  | 
* ---------------------------------------------------  | 
|
1057  | 
* A U B U C |- R  | 
|
1058  | 
*  | 
|
1059  | 
*---------------------------------------------------------------------------*)  | 
|
1060  | 
||
1061  | 
fun DISJ_CASES th1 th2 th3 =  | 
|
1062  | 
let  | 
|
1063  | 
val c = Dcterm.drop_prop (cconcl th1);  | 
|
1064  | 
val (disj1, disj2) = Dcterm.dest_disj c;  | 
|
1065  | 
val th2' = DISCH disj1 th2;  | 
|
1066  | 
val th3' = DISCH disj2 th3;  | 
|
1067  | 
in  | 
|
| 60522 | 1068  | 
    th3' RS (th2' RS (th1 RS @{thm tfl_disjE}))
 | 
| 60520 | 1069  | 
handle THM (msg, _, _) => raise RULES_ERR "DISJ_CASES" msg  | 
1070  | 
end;  | 
|
1071  | 
||
1072  | 
||
1073  | 
(*-----------------------------------------------------------------------------  | 
|
1074  | 
*  | 
|
1075  | 
* |- A1 \/ ... \/ An [A1 |- M, ..., An |- M]  | 
|
1076  | 
* ---------------------------------------------------  | 
|
1077  | 
* |- M  | 
|
1078  | 
*  | 
|
1079  | 
* Note. The list of theorems may be all jumbled up, so we have to  | 
|
1080  | 
* first organize it to align with the first argument (the disjunctive  | 
|
1081  | 
* theorem).  | 
|
1082  | 
*---------------------------------------------------------------------------*)  | 
|
1083  | 
||
1084  | 
fun organize eq = (* a bit slow - analogous to insertion sort *)  | 
|
1085  | 
let fun extract a alist =  | 
|
1086  | 
let fun ex (_,[]) = raise RULES_ERR "organize" "not a permutation.1"  | 
|
1087  | 
| ex(left,h::t) = if (eq h a) then (h,rev left@t) else ex(h::left,t)  | 
|
1088  | 
in ex ([],alist)  | 
|
1089  | 
end  | 
|
1090  | 
fun place [] [] = []  | 
|
1091  | 
| place (a::rst) alist =  | 
|
1092  | 
let val (item,next) = extract a alist  | 
|
1093  | 
in item::place rst next  | 
|
1094  | 
end  | 
|
1095  | 
| place _ _ = raise RULES_ERR "organize" "not a permutation.2"  | 
|
1096  | 
in place  | 
|
1097  | 
end;  | 
|
1098  | 
||
1099  | 
fun DISJ_CASESL disjth thl =  | 
|
1100  | 
let val c = cconcl disjth  | 
|
1101  | 
fun eq th atm =  | 
|
1102  | 
exists (fn t => HOLogic.dest_Trueprop t aconv Thm.term_of atm) (Thm.hyps_of th)  | 
|
1103  | 
val tml = Dcterm.strip_disj c  | 
|
| 60521 | 1104  | 
fun DL _ [] = raise RULES_ERR "DISJ_CASESL" "no cases"  | 
| 60520 | 1105  | 
| DL th [th1] = PROVE_HYP th th1  | 
1106  | 
| DL th [th1,th2] = DISJ_CASES th th1 th2  | 
|
1107  | 
| DL th (th1::rst) =  | 
|
1108  | 
let val tm = #2 (Dcterm.dest_disj (Dcterm.drop_prop(cconcl th)))  | 
|
1109  | 
in DISJ_CASES th th1 (DL (ASSUME tm) rst) end  | 
|
1110  | 
in DL disjth (organize eq tml thl)  | 
|
1111  | 
end;  | 
|
1112  | 
||
1113  | 
||
1114  | 
(*----------------------------------------------------------------------------  | 
|
1115  | 
* Universals  | 
|
1116  | 
*---------------------------------------------------------------------------*)  | 
|
1117  | 
local (* this is fragile *)  | 
|
1118  | 
val prop = Thm.prop_of spec  | 
|
1119  | 
val x = hd (tl (Misc_Legacy.term_vars prop))  | 
|
| 
60642
 
48dd1cefb4ae
simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
 
wenzelm 
parents: 
60524 
diff
changeset
 | 
1120  | 
val TV = dest_TVar (type_of x)  | 
| 69593 | 1121  | 
val gspec = Thm.forall_intr (Thm.cterm_of \<^context> x) spec  | 
| 60520 | 1122  | 
in  | 
1123  | 
fun SPEC tm thm =  | 
|
| 74282 | 1124  | 
let val gspec' =  | 
1125  | 
Drule.instantiate_normalize (TVars.make [(TV, Thm.ctyp_of_cterm tm)], Vars.empty) gspec  | 
|
| 60520 | 1126  | 
in thm RS (Thm.forall_elim tm gspec') end  | 
1127  | 
end;  | 
|
1128  | 
||
1129  | 
fun SPEC_ALL thm = fold SPEC (#1 (Dcterm.strip_forall(cconcl thm))) thm;  | 
|
1130  | 
||
1131  | 
val ISPEC = SPEC  | 
|
1132  | 
val ISPECL = fold ISPEC;  | 
|
1133  | 
||
1134  | 
(* Not optimized! Too complicated. *)  | 
|
1135  | 
local  | 
|
1136  | 
val prop = Thm.prop_of allI  | 
|
1137  | 
val [P] = Misc_Legacy.add_term_vars (prop, [])  | 
|
| 
60642
 
48dd1cefb4ae
simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
 
wenzelm 
parents: 
60524 
diff
changeset
 | 
1138  | 
fun cty_theta ctxt = map (fn (i, (S, ty)) => ((i, S), Thm.ctyp_of ctxt ty))  | 
| 60520 | 1139  | 
fun ctm_theta ctxt =  | 
1140  | 
map (fn (i, (_, tm2)) =>  | 
|
1141  | 
let val ctm2 = Thm.cterm_of ctxt tm2  | 
|
| 
60642
 
48dd1cefb4ae
simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
 
wenzelm 
parents: 
60524 
diff
changeset
 | 
1142  | 
in ((i, Thm.typ_of_cterm ctm2), ctm2) end)  | 
| 60520 | 1143  | 
fun certify ctxt (ty_theta,tm_theta) =  | 
| 74282 | 1144  | 
(TVars.make (cty_theta ctxt (Vartab.dest ty_theta)),  | 
1145  | 
Vars.make (ctm_theta ctxt (Vartab.dest tm_theta)))  | 
|
| 60520 | 1146  | 
in  | 
1147  | 
fun GEN ctxt v th =  | 
|
1148  | 
let val gth = Thm.forall_intr v th  | 
|
1149  | 
val thy = Proof_Context.theory_of ctxt  | 
|
| 69593 | 1150  | 
val Const(\<^const_name>\<open>Pure.all\<close>,_)$Abs(x,ty,rst) = Thm.prop_of gth  | 
| 60520 | 1151  | 
val P' = Abs(x,ty, HOLogic.dest_Trueprop rst) (* get rid of trueprop *)  | 
1152  | 
val theta = Pattern.match thy (P,P') (Vartab.empty, Vartab.empty);  | 
|
1153  | 
val allI2 = Drule.instantiate_normalize (certify ctxt theta) allI  | 
|
1154  | 
val thm = Thm.implies_elim allI2 gth  | 
|
1155  | 
val tp $ (A $ Abs(_,_,M)) = Thm.prop_of thm  | 
|
1156  | 
val prop' = tp $ (A $ Abs(x,ty,M))  | 
|
1157  | 
in ALPHA thm (Thm.cterm_of ctxt prop') end  | 
|
1158  | 
end;  | 
|
1159  | 
||
1160  | 
fun GENL ctxt = fold_rev (GEN ctxt);  | 
|
1161  | 
||
1162  | 
fun GEN_ALL ctxt thm =  | 
|
1163  | 
let  | 
|
1164  | 
val prop = Thm.prop_of thm  | 
|
1165  | 
val vlist = map (Thm.cterm_of ctxt) (Misc_Legacy.add_term_vars (prop, []))  | 
|
1166  | 
in GENL ctxt vlist thm end;  | 
|
1167  | 
||
1168  | 
||
1169  | 
fun MATCH_MP th1 th2 =  | 
|
1170  | 
if (Dcterm.is_forall (Dcterm.drop_prop(cconcl th1)))  | 
|
1171  | 
then MATCH_MP (th1 RS spec) th2  | 
|
1172  | 
else MP th1 th2;  | 
|
1173  | 
||
1174  | 
||
1175  | 
(*----------------------------------------------------------------------------  | 
|
1176  | 
* Existentials  | 
|
1177  | 
*---------------------------------------------------------------------------*)  | 
|
1178  | 
||
1179  | 
||
1180  | 
||
1181  | 
(*---------------------------------------------------------------------------  | 
|
1182  | 
* Existential elimination  | 
|
1183  | 
*  | 
|
1184  | 
* A1 |- ?x.t[x] , A2, "t[v]" |- t'  | 
|
1185  | 
* ------------------------------------ (variable v occurs nowhere)  | 
|
1186  | 
* A1 u A2 |- t'  | 
|
1187  | 
*  | 
|
1188  | 
*---------------------------------------------------------------------------*)  | 
|
1189  | 
||
1190  | 
fun CHOOSE ctxt (fvar, exth) fact =  | 
|
1191  | 
let  | 
|
1192  | 
val lam = #2 (Dcterm.dest_comb (Dcterm.drop_prop (cconcl exth)))  | 
|
1193  | 
val redex = Dcterm.capply lam fvar  | 
|
1194  | 
val t$u = Thm.term_of redex  | 
|
1195  | 
val residue = Thm.cterm_of ctxt (Term.betapply (t, u))  | 
|
1196  | 
in  | 
|
| 60522 | 1197  | 
    GEN ctxt fvar (DISCH residue fact) RS (exth RS @{thm tfl_exE})
 | 
| 60520 | 1198  | 
handle THM (msg, _, _) => raise RULES_ERR "CHOOSE" msg  | 
1199  | 
end;  | 
|
1200  | 
||
| 60781 | 1201  | 
fun EXISTS ctxt (template,witness) thm =  | 
| 60520 | 1202  | 
let val abstr = #2 (Dcterm.dest_comb template) in  | 
| 60781 | 1203  | 
    thm RS (infer_instantiate ctxt [(("P", 0), abstr), (("x", 0), witness)] exI)
 | 
| 60520 | 1204  | 
handle THM (msg, _, _) => raise RULES_ERR "EXISTS" msg  | 
| 60781 | 1205  | 
end;  | 
| 60520 | 1206  | 
|
1207  | 
(*----------------------------------------------------------------------------  | 
|
1208  | 
*  | 
|
1209  | 
* A |- M[x_1,...,x_n]  | 
|
1210  | 
* ---------------------------- [(x |-> y)_1,...,(x |-> y)_n]  | 
|
1211  | 
* A |- ?y_1...y_n. M  | 
|
1212  | 
*  | 
|
1213  | 
*---------------------------------------------------------------------------*)  | 
|
1214  | 
(* Could be improved, but needs "subst_free" for certified terms *)  | 
|
1215  | 
||
1216  | 
fun IT_EXISTS ctxt blist th =  | 
|
1217  | 
let  | 
|
1218  | 
val blist' = map (apply2 Thm.term_of) blist  | 
|
1219  | 
    fun ex v M = Thm.cterm_of ctxt (USyntax.mk_exists{Bvar=v,Body = M})
 | 
|
1220  | 
in  | 
|
1221  | 
fold_rev (fn (b as (r1,r2)) => fn thm =>  | 
|
| 60781 | 1222  | 
EXISTS ctxt (ex r2 (subst_free [b]  | 
| 60520 | 1223  | 
(HOLogic.dest_Trueprop(Thm.prop_of thm))), Thm.cterm_of ctxt r1)  | 
1224  | 
thm)  | 
|
1225  | 
blist' th  | 
|
1226  | 
end;  | 
|
1227  | 
||
1228  | 
(*----------------------------------------------------------------------------  | 
|
1229  | 
* Rewriting  | 
|
1230  | 
*---------------------------------------------------------------------------*)  | 
|
1231  | 
||
1232  | 
fun SUBS ctxt thl =  | 
|
1233  | 
rewrite_rule ctxt (map (fn th => th RS eq_reflection handle THM _ => th) thl);  | 
|
1234  | 
||
1235  | 
val rew_conv = Raw_Simplifier.rewrite_cterm (true, false, false) (K (K NONE));  | 
|
1236  | 
||
1237  | 
fun simpl_conv ctxt thl ctm =  | 
|
| 
67710
 
cc2db3239932
added HOLogic.mk_obj_eq convenience and eliminated some clones;
 
wenzelm 
parents: 
66251 
diff
changeset
 | 
1238  | 
HOLogic.mk_obj_eq (rew_conv (ctxt addsimps thl) ctm);  | 
| 60520 | 1239  | 
|
1240  | 
||
| 60522 | 1241  | 
fun RIGHT_ASSOC ctxt = rewrite_rule ctxt @{thms tfl_disj_assoc};
 | 
| 60520 | 1242  | 
|
1243  | 
||
1244  | 
||
1245  | 
(*---------------------------------------------------------------------------  | 
|
1246  | 
* TERMINATION CONDITION EXTRACTION  | 
|
1247  | 
*---------------------------------------------------------------------------*)  | 
|
1248  | 
||
1249  | 
||
1250  | 
(* Object language quantifier, i.e., "!" *)  | 
|
1251  | 
fun Forall v M = USyntax.mk_forall{Bvar=v, Body=M};
 | 
|
1252  | 
||
1253  | 
||
1254  | 
(* Fragile: it's a cong if it is not "R y x ==> cut f R x y = f y" *)  | 
|
1255  | 
fun is_cong thm =  | 
|
1256  | 
case (Thm.prop_of thm) of  | 
|
| 69593 | 1257  | 
(Const(\<^const_name>\<open>Pure.imp\<close>,_)$(Const(\<^const_name>\<open>Trueprop\<close>,_)$ _) $  | 
1258  | 
(Const(\<^const_name>\<open>Pure.eq\<close>,_) $ (Const (\<^const_name>\<open>Wfrec.cut\<close>,_) $ _ $ _ $ _ $ _) $ _)) =>  | 
|
| 60520 | 1259  | 
false  | 
1260  | 
| _ => true;  | 
|
1261  | 
||
1262  | 
||
| 69593 | 1263  | 
fun dest_equal(Const (\<^const_name>\<open>Pure.eq\<close>,_) $  | 
1264  | 
(Const (\<^const_name>\<open>Trueprop\<close>,_) $ lhs)  | 
|
1265  | 
               $ (Const (\<^const_name>\<open>Trueprop\<close>,_) $ rhs)) = {lhs=lhs, rhs=rhs}
 | 
|
1266  | 
  | dest_equal(Const (\<^const_name>\<open>Pure.eq\<close>,_) $ lhs $ rhs) = {lhs=lhs, rhs=rhs}
 | 
|
| 60520 | 1267  | 
| dest_equal tm = USyntax.dest_eq tm;  | 
1268  | 
||
1269  | 
fun get_lhs tm = #lhs(dest_equal (HOLogic.dest_Trueprop tm));  | 
|
1270  | 
||
| 69593 | 1271  | 
fun dest_all used (Const(\<^const_name>\<open>Pure.all\<close>,_) $ (a as Abs _)) = USyntax.dest_abs used a  | 
| 60520 | 1272  | 
| dest_all _ _ = raise RULES_ERR "dest_all" "not a !!";  | 
1273  | 
||
1274  | 
val is_all = can (dest_all []);  | 
|
1275  | 
||
1276  | 
fun strip_all used fm =  | 
|
1277  | 
if (is_all fm)  | 
|
1278  | 
   then let val ({Bvar, Body}, used') = dest_all used fm
 | 
|
1279  | 
val (bvs, core, used'') = strip_all used' Body  | 
|
1280  | 
in ((Bvar::bvs), core, used'')  | 
|
1281  | 
end  | 
|
1282  | 
else ([], fm, used);  | 
|
1283  | 
||
| 69593 | 1284  | 
fun list_break_all(Const(\<^const_name>\<open>Pure.all\<close>,_) $ Abs (s,ty,body)) =  | 
| 60520 | 1285  | 
let val (L,core) = list_break_all body  | 
1286  | 
in ((s,ty)::L, core)  | 
|
1287  | 
end  | 
|
1288  | 
| list_break_all tm = ([],tm);  | 
|
1289  | 
||
1290  | 
(*---------------------------------------------------------------------------  | 
|
1291  | 
* Rename a term of the form  | 
|
1292  | 
*  | 
|
1293  | 
* !!x1 ...xn. x1=M1 ==> ... ==> xn=Mn  | 
|
1294  | 
* ==> ((%v1...vn. Q) x1 ... xn = g x1 ... xn.  | 
|
1295  | 
* to one of  | 
|
1296  | 
*  | 
|
1297  | 
* !!v1 ... vn. v1=M1 ==> ... ==> vn=Mn  | 
|
1298  | 
* ==> ((%v1...vn. Q) v1 ... vn = g v1 ... vn.  | 
|
1299  | 
*  | 
|
1300  | 
* This prevents name problems in extraction, and helps the result to read  | 
|
1301  | 
* better. There is a problem with varstructs, since they can introduce more  | 
|
1302  | 
* than n variables, and some extra reasoning needs to be done.  | 
|
1303  | 
*---------------------------------------------------------------------------*)  | 
|
1304  | 
||
1305  | 
fun get ([],_,L) = rev L  | 
|
1306  | 
| get (ant::rst,n,L) =  | 
|
1307  | 
case (list_break_all ant)  | 
|
1308  | 
of ([],_) => get (rst, n+1,L)  | 
|
| 60521 | 1309  | 
| (_,body) =>  | 
| 60520 | 1310  | 
let val eq = Logic.strip_imp_concl body  | 
| 60521 | 1311  | 
val (f,_) = USyntax.strip_comb (get_lhs eq)  | 
| 60520 | 1312  | 
val (vstrl,_) = USyntax.strip_abs f  | 
1313  | 
val names =  | 
|
1314  | 
Name.variant_list (Misc_Legacy.add_term_names(body, [])) (map (#1 o dest_Free) vstrl)  | 
|
1315  | 
in get (rst, n+1, (names,n)::L) end  | 
|
1316  | 
handle TERM _ => get (rst, n+1, L)  | 
|
1317  | 
| Utils.ERR _ => get (rst, n+1, L);  | 
|
1318  | 
||
1319  | 
(* Note: Thm.rename_params_rule counts from 1, not 0 *)  | 
|
1320  | 
fun rename thm =  | 
|
1321  | 
let  | 
|
1322  | 
val ants = Logic.strip_imp_prems (Thm.prop_of thm)  | 
|
1323  | 
val news = get (ants,1,[])  | 
|
1324  | 
in fold Thm.rename_params_rule news thm end;  | 
|
1325  | 
||
1326  | 
||
1327  | 
(*---------------------------------------------------------------------------  | 
|
1328  | 
* Beta-conversion to the rhs of an equation (taken from hol90/drule.sml)  | 
|
1329  | 
*---------------------------------------------------------------------------*)  | 
|
1330  | 
||
1331  | 
fun list_beta_conv tm =  | 
|
1332  | 
let fun rbeta th = Thm.transitive th (Thm.beta_conversion false (#2(Dcterm.dest_eq(cconcl th))))  | 
|
1333  | 
fun iter [] = Thm.reflexive tm  | 
|
1334  | 
| iter (v::rst) = rbeta (Thm.combination(iter rst) (Thm.reflexive v))  | 
|
1335  | 
in iter end;  | 
|
1336  | 
||
1337  | 
||
1338  | 
(*---------------------------------------------------------------------------  | 
|
1339  | 
* Trace information for the rewriter  | 
|
1340  | 
*---------------------------------------------------------------------------*)  | 
|
1341  | 
val tracing = Unsynchronized.ref false;  | 
|
1342  | 
||
1343  | 
fun say s = if !tracing then writeln s else ();  | 
|
1344  | 
||
1345  | 
fun print_thms ctxt s L =  | 
|
| 61268 | 1346  | 
say (cat_lines (s :: map (Thm.string_of_thm ctxt) L));  | 
| 60520 | 1347  | 
|
1348  | 
fun print_term ctxt s t =  | 
|
1349  | 
say (cat_lines [s, Syntax.string_of_term ctxt t]);  | 
|
1350  | 
||
1351  | 
||
1352  | 
(*---------------------------------------------------------------------------  | 
|
1353  | 
* General abstraction handlers, should probably go in USyntax.  | 
|
1354  | 
*---------------------------------------------------------------------------*)  | 
|
1355  | 
fun mk_aabs (vstr, body) =  | 
|
1356  | 
  USyntax.mk_abs {Bvar = vstr, Body = body}
 | 
|
1357  | 
  handle Utils.ERR _ => USyntax.mk_pabs {varstruct = vstr, body = body};
 | 
|
1358  | 
||
1359  | 
fun list_mk_aabs (vstrl,tm) =  | 
|
1360  | 
fold_rev (fn vstr => fn tm => mk_aabs(vstr,tm)) vstrl tm;  | 
|
1361  | 
||
1362  | 
fun dest_aabs used tm =  | 
|
1363  | 
   let val ({Bvar,Body}, used') = USyntax.dest_abs used tm
 | 
|
1364  | 
in (Bvar, Body, used') end  | 
|
1365  | 
handle Utils.ERR _ =>  | 
|
1366  | 
     let val {varstruct, body, used} = USyntax.dest_pabs used tm
 | 
|
1367  | 
in (varstruct, body, used) end;  | 
|
1368  | 
||
1369  | 
fun strip_aabs used tm =  | 
|
1370  | 
let val (vstr, body, used') = dest_aabs used tm  | 
|
1371  | 
val (bvs, core, used'') = strip_aabs used' body  | 
|
1372  | 
in (vstr::bvs, core, used'') end  | 
|
1373  | 
handle Utils.ERR _ => ([], tm, used);  | 
|
1374  | 
||
1375  | 
fun dest_combn tm 0 = (tm,[])  | 
|
1376  | 
| dest_combn tm n =  | 
|
1377  | 
     let val {Rator,Rand} = USyntax.dest_comb tm
 | 
|
1378  | 
val (f,rands) = dest_combn Rator (n-1)  | 
|
1379  | 
in (f,Rand::rands)  | 
|
1380  | 
end;  | 
|
1381  | 
||
1382  | 
||
1383  | 
||
1384  | 
||
1385  | 
local fun dest_pair M = let val {fst,snd} = USyntax.dest_pair M in (fst,snd) end
 | 
|
1386  | 
fun mk_fst tm =  | 
|
| 69593 | 1387  | 
let val ty as Type(\<^type_name>\<open>Product_Type.prod\<close>, [fty,sty]) = type_of tm  | 
1388  | 
in Const (\<^const_name>\<open>Product_Type.fst\<close>, ty --> fty) $ tm end  | 
|
| 60520 | 1389  | 
fun mk_snd tm =  | 
| 69593 | 1390  | 
let val ty as Type(\<^type_name>\<open>Product_Type.prod\<close>, [fty,sty]) = type_of tm  | 
1391  | 
in Const (\<^const_name>\<open>Product_Type.snd\<close>, ty --> sty) $ tm end  | 
|
| 60520 | 1392  | 
in  | 
1393  | 
fun XFILL tych x vstruct =  | 
|
1394  | 
let fun traverse p xocc L =  | 
|
1395  | 
if (is_Free p)  | 
|
1396  | 
then tych xocc::L  | 
|
1397  | 
else let val (p1,p2) = dest_pair p  | 
|
1398  | 
in traverse p1 (mk_fst xocc) (traverse p2 (mk_snd xocc) L)  | 
|
1399  | 
end  | 
|
1400  | 
in  | 
|
1401  | 
traverse vstruct x []  | 
|
1402  | 
end end;  | 
|
1403  | 
||
1404  | 
(*---------------------------------------------------------------------------  | 
|
1405  | 
* Replace a free tuple (vstr) by a universally quantified variable (a).  | 
|
1406  | 
* Note that the notion of "freeness" for a tuple is different than for a  | 
|
1407  | 
* variable: if variables in the tuple also occur in any other place than  | 
|
1408  | 
* an occurrences of the tuple, they aren't "free" (which is thus probably  | 
|
1409  | 
* the wrong word to use).  | 
|
1410  | 
*---------------------------------------------------------------------------*)  | 
|
1411  | 
||
1412  | 
fun VSTRUCT_ELIM ctxt tych a vstr th =  | 
|
1413  | 
let val L = USyntax.free_vars_lr vstr  | 
|
1414  | 
val bind1 = tych (HOLogic.mk_Trueprop (HOLogic.mk_eq(a,vstr)))  | 
|
1415  | 
val thm1 = Thm.implies_intr bind1 (SUBS ctxt [SYM(Thm.assume bind1)] th)  | 
|
1416  | 
val thm2 = forall_intr_list (map tych L) thm1  | 
|
1417  | 
val thm3 = forall_elim_list (XFILL tych a vstr) thm2  | 
|
1418  | 
in refl RS  | 
|
1419  | 
     rewrite_rule ctxt [Thm.symmetric (@{thm surjective_pairing} RS eq_reflection)] thm3
 | 
|
1420  | 
end;  | 
|
1421  | 
||
1422  | 
fun PGEN ctxt tych a vstr th =  | 
|
1423  | 
let val a1 = tych a  | 
|
| 60781 | 1424  | 
in Thm.forall_intr a1 (VSTRUCT_ELIM ctxt tych a vstr th) end;  | 
| 60520 | 1425  | 
|
1426  | 
||
1427  | 
(*---------------------------------------------------------------------------  | 
|
1428  | 
* Takes apart a paired beta-redex, looking like "(\(x,y).N) vstr", into  | 
|
1429  | 
*  | 
|
1430  | 
* (([x,y],N),vstr)  | 
|
1431  | 
*---------------------------------------------------------------------------*)  | 
|
1432  | 
fun dest_pbeta_redex used M n =  | 
|
1433  | 
let val (f,args) = dest_combn M n  | 
|
| 60521 | 1434  | 
val _ = dest_aabs used f  | 
| 60520 | 1435  | 
in (strip_aabs used f,args)  | 
1436  | 
end;  | 
|
1437  | 
||
1438  | 
fun pbeta_redex M n = can (fn t => dest_pbeta_redex [] t n) M;  | 
|
1439  | 
||
1440  | 
fun dest_impl tm =  | 
|
1441  | 
let val ants = Logic.strip_imp_prems tm  | 
|
1442  | 
val eq = Logic.strip_imp_concl tm  | 
|
1443  | 
in (ants,get_lhs eq)  | 
|
1444  | 
end;  | 
|
1445  | 
||
1446  | 
fun restricted t = is_some (USyntax.find_term  | 
|
| 69593 | 1447  | 
(fn (Const(\<^const_name>\<open>Wfrec.cut\<close>,_)) =>true | _ => false)  | 
| 60520 | 1448  | 
t)  | 
1449  | 
||
1450  | 
fun CONTEXT_REWRITE_RULE main_ctxt (func, G, cut_lemma, congs) th =  | 
|
1451  | 
let val globals = func::G  | 
|
1452  | 
val ctxt0 = empty_simpset main_ctxt  | 
|
1453  | 
     val pbeta_reduce = simpl_conv ctxt0 [@{thm split_conv} RS eq_reflection];
 | 
|
1454  | 
val tc_list = Unsynchronized.ref []: term list Unsynchronized.ref  | 
|
1455  | 
val cut_lemma' = cut_lemma RS eq_reflection  | 
|
1456  | 
fun prover used ctxt thm =  | 
|
1457  | 
let fun cong_prover ctxt thm =  | 
|
| 60521 | 1458  | 
let val _ = say "cong_prover:"  | 
| 60520 | 1459  | 
val cntxt = Simplifier.prems_of ctxt  | 
| 60521 | 1460  | 
val _ = print_thms ctxt "cntxt:" cntxt  | 
1461  | 
val _ = say "cong rule:"  | 
|
| 61268 | 1462  | 
val _ = say (Thm.string_of_thm ctxt thm)  | 
| 60520 | 1463  | 
(* Unquantified eliminate *)  | 
1464  | 
fun uq_eliminate (thm,imp) =  | 
|
1465  | 
let val tych = Thm.cterm_of ctxt  | 
|
1466  | 
val _ = print_term ctxt "To eliminate:" imp  | 
|
1467  | 
val ants = map tych (Logic.strip_imp_prems imp)  | 
|
1468  | 
val eq = Logic.strip_imp_concl imp  | 
|
1469  | 
val lhs = tych(get_lhs eq)  | 
|
1470  | 
val ctxt' = Simplifier.add_prems (map ASSUME ants) ctxt  | 
|
1471  | 
val lhs_eq_lhs1 = Raw_Simplifier.rewrite_cterm (false,true,false) (prover used) ctxt' lhs  | 
|
1472  | 
handle Utils.ERR _ => Thm.reflexive lhs  | 
|
1473  | 
val _ = print_thms ctxt' "proven:" [lhs_eq_lhs1]  | 
|
1474  | 
val lhs_eq_lhs2 = implies_intr_list ants lhs_eq_lhs1  | 
|
| 
67710
 
cc2db3239932
added HOLogic.mk_obj_eq convenience and eliminated some clones;
 
wenzelm 
parents: 
66251 
diff
changeset
 | 
1475  | 
val lhs_eeq_lhs2 = HOLogic.mk_obj_eq lhs_eq_lhs2  | 
| 60520 | 1476  | 
in  | 
1477  | 
lhs_eeq_lhs2 COMP thm  | 
|
1478  | 
end  | 
|
1479  | 
fun pq_eliminate (thm, vlist, imp_body, lhs_eq) =  | 
|
1480  | 
let val ((vstrl, _, used'), args) = dest_pbeta_redex used lhs_eq (length vlist)  | 
|
| 60521 | 1481  | 
val _ = forall (op aconv) (ListPair.zip (vlist, args))  | 
| 60520 | 1482  | 
orelse error "assertion failed in CONTEXT_REWRITE_RULE"  | 
1483  | 
val imp_body1 = subst_free (ListPair.zip (args, vstrl))  | 
|
1484  | 
imp_body  | 
|
1485  | 
val tych = Thm.cterm_of ctxt  | 
|
1486  | 
val ants1 = map tych (Logic.strip_imp_prems imp_body1)  | 
|
1487  | 
val eq1 = Logic.strip_imp_concl imp_body1  | 
|
1488  | 
val Q = get_lhs eq1  | 
|
1489  | 
val QeqQ1 = pbeta_reduce (tych Q)  | 
|
1490  | 
val Q1 = #2(Dcterm.dest_eq(cconcl QeqQ1))  | 
|
1491  | 
val ctxt' = Simplifier.add_prems (map ASSUME ants1) ctxt  | 
|
1492  | 
val Q1eeqQ2 = Raw_Simplifier.rewrite_cterm (false,true,false) (prover used') ctxt' Q1  | 
|
1493  | 
handle Utils.ERR _ => Thm.reflexive Q1  | 
|
1494  | 
val Q2 = #2 (Logic.dest_equals (Thm.prop_of Q1eeqQ2))  | 
|
1495  | 
val Q3 = tych(list_comb(list_mk_aabs(vstrl,Q2),vstrl))  | 
|
1496  | 
val Q2eeqQ3 = Thm.symmetric(pbeta_reduce Q3 RS eq_reflection)  | 
|
1497  | 
val thA = Thm.transitive(QeqQ1 RS eq_reflection) Q1eeqQ2  | 
|
1498  | 
val QeeqQ3 = Thm.transitive thA Q2eeqQ3 handle THM _ =>  | 
|
| 
67710
 
cc2db3239932
added HOLogic.mk_obj_eq convenience and eliminated some clones;
 
wenzelm 
parents: 
66251 
diff
changeset
 | 
1499  | 
(HOLogic.mk_obj_eq Q2eeqQ3  | 
| 
 
cc2db3239932
added HOLogic.mk_obj_eq convenience and eliminated some clones;
 
wenzelm 
parents: 
66251 
diff
changeset
 | 
1500  | 
RS (HOLogic.mk_obj_eq thA RS trans))  | 
| 60520 | 1501  | 
RS eq_reflection  | 
1502  | 
val impth = implies_intr_list ants1 QeeqQ3  | 
|
| 
67710
 
cc2db3239932
added HOLogic.mk_obj_eq convenience and eliminated some clones;
 
wenzelm 
parents: 
66251 
diff
changeset
 | 
1503  | 
val impth1 = HOLogic.mk_obj_eq impth  | 
| 60520 | 1504  | 
(* Need to abstract *)  | 
1505  | 
val ant_th = Utils.itlist2 (PGEN ctxt' tych) args vstrl impth1  | 
|
1506  | 
in ant_th COMP thm  | 
|
1507  | 
end  | 
|
1508  | 
fun q_eliminate (thm, imp) =  | 
|
1509  | 
let val (vlist, imp_body, used') = strip_all used imp  | 
|
1510  | 
val (ants,Q) = dest_impl imp_body  | 
|
1511  | 
in if (pbeta_redex Q) (length vlist)  | 
|
1512  | 
then pq_eliminate (thm, vlist, imp_body, Q)  | 
|
1513  | 
else  | 
|
1514  | 
let val tych = Thm.cterm_of ctxt  | 
|
1515  | 
val ants1 = map tych ants  | 
|
1516  | 
val ctxt' = Simplifier.add_prems (map ASSUME ants1) ctxt  | 
|
1517  | 
val Q_eeq_Q1 = Raw_Simplifier.rewrite_cterm  | 
|
1518  | 
(false,true,false) (prover used') ctxt' (tych Q)  | 
|
1519  | 
handle Utils.ERR _ => Thm.reflexive (tych Q)  | 
|
1520  | 
val lhs_eeq_lhs2 = implies_intr_list ants1 Q_eeq_Q1  | 
|
| 
67710
 
cc2db3239932
added HOLogic.mk_obj_eq convenience and eliminated some clones;
 
wenzelm 
parents: 
66251 
diff
changeset
 | 
1521  | 
val lhs_eq_lhs2 = HOLogic.mk_obj_eq lhs_eeq_lhs2  | 
| 60520 | 1522  | 
val ant_th = forall_intr_list(map tych vlist)lhs_eq_lhs2  | 
1523  | 
in  | 
|
1524  | 
ant_th COMP thm  | 
|
1525  | 
end end  | 
|
1526  | 
||
1527  | 
fun eliminate thm =  | 
|
1528  | 
case Thm.prop_of thm of  | 
|
| 69593 | 1529  | 
Const(\<^const_name>\<open>Pure.imp\<close>,_) $ imp $ _ =>  | 
| 60520 | 1530  | 
eliminate  | 
1531  | 
(if not(is_all imp)  | 
|
1532  | 
then uq_eliminate (thm, imp)  | 
|
1533  | 
else q_eliminate (thm, imp))  | 
|
1534  | 
(* Assume that the leading constant is ==, *)  | 
|
1535  | 
| _ => thm (* if it is not a ==> *)  | 
|
1536  | 
in SOME(eliminate (rename thm)) end  | 
|
1537  | 
handle Utils.ERR _ => NONE (* FIXME handle THM as well?? *)  | 
|
1538  | 
||
1539  | 
fun restrict_prover ctxt thm =  | 
|
1540  | 
let val _ = say "restrict_prover:"  | 
|
1541  | 
val cntxt = rev (Simplifier.prems_of ctxt)  | 
|
1542  | 
val _ = print_thms ctxt "cntxt:" cntxt  | 
|
| 69593 | 1543  | 
val Const(\<^const_name>\<open>Pure.imp\<close>,_) $ (Const(\<^const_name>\<open>Trueprop\<close>,_) $ A) $ _ =  | 
| 60520 | 1544  | 
Thm.prop_of thm  | 
1545  | 
fun genl tm = let val vlist = subtract (op aconv) globals  | 
|
1546  | 
(Misc_Legacy.add_term_frees(tm,[]))  | 
|
1547  | 
in fold_rev Forall vlist tm  | 
|
1548  | 
end  | 
|
1549  | 
(*--------------------------------------------------------------  | 
|
1550  | 
* This actually isn't quite right, since it will think that  | 
|
1551  | 
* not-fully applied occs. of "f" in the context mean that the  | 
|
1552  | 
* current call is nested. The real solution is to pass in a  | 
|
1553  | 
* term "f v1..vn" which is a pattern that any full application  | 
|
1554  | 
* of "f" will match.  | 
|
1555  | 
*-------------------------------------------------------------*)  | 
|
1556  | 
val func_name = #1(dest_Const func)  | 
|
1557  | 
fun is_func (Const (name,_)) = (name = func_name)  | 
|
1558  | 
| is_func _ = false  | 
|
1559  | 
val rcontext = rev cntxt  | 
|
1560  | 
val cncl = HOLogic.dest_Trueprop o Thm.prop_of  | 
|
1561  | 
val antl = case rcontext of [] => []  | 
|
1562  | 
| _ => [USyntax.list_mk_conj(map cncl rcontext)]  | 
|
1563  | 
val TC = genl(USyntax.list_mk_imp(antl, A))  | 
|
1564  | 
val _ = print_term ctxt "func:" func  | 
|
1565  | 
val _ = print_term ctxt "TC:" (HOLogic.mk_Trueprop TC)  | 
|
1566  | 
val _ = tc_list := (TC :: !tc_list)  | 
|
1567  | 
val nestedp = is_some (USyntax.find_term is_func TC)  | 
|
1568  | 
val _ = if nestedp then say "nested" else say "not_nested"  | 
|
1569  | 
val th' = if nestedp then raise RULES_ERR "solver" "nested function"  | 
|
1570  | 
else let val cTC = Thm.cterm_of ctxt (HOLogic.mk_Trueprop TC)  | 
|
1571  | 
in case rcontext of  | 
|
1572  | 
[] => SPEC_ALL(ASSUME cTC)  | 
|
1573  | 
| _ => MP (SPEC_ALL (ASSUME cTC))  | 
|
1574  | 
(LIST_CONJ rcontext)  | 
|
1575  | 
end  | 
|
1576  | 
val th'' = th' RS thm  | 
|
1577  | 
in SOME (th'')  | 
|
1578  | 
end handle Utils.ERR _ => NONE (* FIXME handle THM as well?? *)  | 
|
1579  | 
in  | 
|
1580  | 
(if (is_cong thm) then cong_prover else restrict_prover) ctxt thm  | 
|
1581  | 
end  | 
|
1582  | 
val ctm = Thm.cprop_of th  | 
|
1583  | 
val names = Misc_Legacy.add_term_names (Thm.term_of ctm, [])  | 
|
1584  | 
val th1 =  | 
|
1585  | 
Raw_Simplifier.rewrite_cterm (false, true, false)  | 
|
1586  | 
(prover names) (ctxt0 addsimps [cut_lemma'] |> fold Simplifier.add_eqcong congs) ctm  | 
|
1587  | 
val th2 = Thm.equal_elim th1 th  | 
|
1588  | 
in  | 
|
1589  | 
(th2, filter_out restricted (!tc_list))  | 
|
1590  | 
end;  | 
|
1591  | 
||
1592  | 
||
1593  | 
fun prove ctxt strict (t, tac) =  | 
|
1594  | 
let  | 
|
| 70308 | 1595  | 
val ctxt' = Proof_Context.augment t ctxt;  | 
| 60520 | 1596  | 
in  | 
1597  | 
if strict  | 
|
1598  | 
then Goal.prove ctxt' [] [] t (K tac)  | 
|
1599  | 
else Goal.prove ctxt' [] [] t (K tac)  | 
|
1600  | 
handle ERROR msg => (warning msg; raise RULES_ERR "prove" msg)  | 
|
1601  | 
end;  | 
|
1602  | 
||
1603  | 
end;  | 
|
1604  | 
||
1605  | 
||
| 60521 | 1606  | 
|
| 60520 | 1607  | 
(*** theory operations ***)  | 
1608  | 
||
1609  | 
structure Thry: THRY =  | 
|
1610  | 
struct  | 
|
1611  | 
||
1612  | 
||
1613  | 
fun THRY_ERR func mesg = Utils.ERR {module = "Thry", func = func, mesg = mesg};
 | 
|
1614  | 
||
1615  | 
||
1616  | 
(*---------------------------------------------------------------------------  | 
|
1617  | 
* Matching  | 
|
1618  | 
*---------------------------------------------------------------------------*)  | 
|
1619  | 
||
1620  | 
local  | 
|
1621  | 
||
1622  | 
fun tybind (ixn, (S, T)) = (TVar (ixn, S), T);  | 
|
1623  | 
||
1624  | 
in  | 
|
1625  | 
||
1626  | 
fun match_term thry pat ob =  | 
|
1627  | 
let  | 
|
1628  | 
val (ty_theta, tm_theta) = Pattern.match thry (pat,ob) (Vartab.empty, Vartab.empty);  | 
|
1629  | 
fun tmbind (ixn, (T, t)) = (Var (ixn, Envir.subst_type ty_theta T), t)  | 
|
1630  | 
in (map tmbind (Vartab.dest tm_theta), map tybind (Vartab.dest ty_theta))  | 
|
1631  | 
end;  | 
|
1632  | 
||
1633  | 
fun match_type thry pat ob =  | 
|
1634  | 
map tybind (Vartab.dest (Sign.typ_match thry (pat, ob) Vartab.empty));  | 
|
1635  | 
||
1636  | 
end;  | 
|
1637  | 
||
1638  | 
||
1639  | 
(*---------------------------------------------------------------------------  | 
|
1640  | 
* Typing  | 
|
1641  | 
*---------------------------------------------------------------------------*)  | 
|
1642  | 
||
1643  | 
fun typecheck thy t =  | 
|
1644  | 
Thm.global_cterm_of thy t  | 
|
1645  | 
handle TYPE (msg, _, _) => raise THRY_ERR "typecheck" msg  | 
|
1646  | 
| TERM (msg, _) => raise THRY_ERR "typecheck" msg;  | 
|
1647  | 
||
1648  | 
||
1649  | 
(*---------------------------------------------------------------------------  | 
|
1650  | 
* Get information about datatypes  | 
|
1651  | 
*---------------------------------------------------------------------------*)  | 
|
1652  | 
||
1653  | 
fun match_info thy dtco =  | 
|
1654  | 
case (BNF_LFP_Compat.get_info thy [BNF_LFP_Compat.Keep_Nesting] dtco,  | 
|
1655  | 
BNF_LFP_Compat.get_constrs thy dtco) of  | 
|
1656  | 
      (SOME {case_name, ... }, SOME constructors) =>
 | 
|
1657  | 
        SOME {case_const = Const (case_name, Sign.the_const_type thy case_name), constructors = map Const constructors}
 | 
|
1658  | 
| _ => NONE;  | 
|
1659  | 
||
1660  | 
fun induct_info thy dtco = case BNF_LFP_Compat.get_info thy [BNF_LFP_Compat.Keep_Nesting] dtco of  | 
|
1661  | 
NONE => NONE  | 
|
1662  | 
      | SOME {nchotomy, ...} =>
 | 
|
1663  | 
          SOME {nchotomy = nchotomy,
 | 
|
1664  | 
constructors = (map Const o the o BNF_LFP_Compat.get_constrs thy) dtco};  | 
|
1665  | 
||
1666  | 
fun extract_info thy =  | 
|
1667  | 
let val infos = map snd (Symtab.dest (BNF_LFP_Compat.get_all thy [BNF_LFP_Compat.Keep_Nesting]))  | 
|
1668  | 
 in {case_congs = map (mk_meta_eq o #case_cong) infos,
 | 
|
1669  | 
case_rewrites = maps (map mk_meta_eq o #case_rewrites) infos}  | 
|
1670  | 
end;  | 
|
1671  | 
||
1672  | 
||
1673  | 
end;  | 
|
1674  | 
||
1675  | 
||
| 60521 | 1676  | 
|
| 60520 | 1677  | 
(*** first part of main module ***)  | 
1678  | 
||
1679  | 
structure Prim: PRIM =  | 
|
1680  | 
struct  | 
|
1681  | 
||
1682  | 
val trace = Unsynchronized.ref false;  | 
|
1683  | 
||
1684  | 
||
1685  | 
fun TFL_ERR func mesg = Utils.ERR {module = "Tfl", func = func, mesg = mesg};
 | 
|
1686  | 
||
1687  | 
val concl = #2 o Rules.dest_thm;  | 
|
1688  | 
||
1689  | 
val list_mk_type = Utils.end_itlist (curry (op -->));  | 
|
1690  | 
||
1691  | 
fun front_last [] = raise TFL_ERR "front_last" "empty list"  | 
|
1692  | 
| front_last [x] = ([],x)  | 
|
1693  | 
| front_last (h::t) =  | 
|
1694  | 
let val (pref,x) = front_last t  | 
|
1695  | 
in  | 
|
1696  | 
(h::pref,x)  | 
|
1697  | 
end;  | 
|
1698  | 
||
1699  | 
||
1700  | 
(*---------------------------------------------------------------------------  | 
|
1701  | 
* The next function is common to pattern-match translation and  | 
|
1702  | 
* proof of completeness of cases for the induction theorem.  | 
|
1703  | 
*  | 
|
1704  | 
* The curried function "gvvariant" returns a function to generate distinct  | 
|
1705  | 
* variables that are guaranteed not to be in names. The names of  | 
|
1706  | 
* the variables go u, v, ..., z, aa, ..., az, ... The returned  | 
|
1707  | 
* function contains embedded refs!  | 
|
1708  | 
*---------------------------------------------------------------------------*)  | 
|
1709  | 
fun gvvariant names =  | 
|
1710  | 
let val slist = Unsynchronized.ref names  | 
|
1711  | 
val vname = Unsynchronized.ref "u"  | 
|
1712  | 
fun new() =  | 
|
1713  | 
if member (op =) (!slist) (!vname)  | 
|
1714  | 
then (vname := Symbol.bump_string (!vname); new())  | 
|
1715  | 
else (slist := !vname :: !slist; !vname)  | 
|
1716  | 
in  | 
|
1717  | 
fn ty => Free(new(), ty)  | 
|
1718  | 
end;  | 
|
1719  | 
||
1720  | 
||
1721  | 
(*---------------------------------------------------------------------------  | 
|
1722  | 
* Used in induction theorem production. This is the simple case of  | 
|
1723  | 
* partitioning up pattern rows by the leading constructor.  | 
|
1724  | 
*---------------------------------------------------------------------------*)  | 
|
1725  | 
fun ipartition gv (constructors,rows) =  | 
|
1726  | 
let fun pfail s = raise TFL_ERR "partition.part" s  | 
|
1727  | 
      fun part {constrs = [],   rows = [],   A} = rev A
 | 
|
1728  | 
        | part {constrs = [],   rows = _::_, A} = pfail"extra cases in defn"
 | 
|
1729  | 
        | part {constrs = _::_, rows = [],   A} = pfail"cases missing in defn"
 | 
|
1730  | 
        | part {constrs = c::crst, rows,     A} =
 | 
|
1731  | 
let val (c, T) = dest_Const c  | 
|
1732  | 
val L = binder_types T  | 
|
1733  | 
val (in_group, not_in_group) =  | 
|
1734  | 
fold_rev (fn (row as (p::rst, rhs)) =>  | 
|
1735  | 
fn (in_group,not_in_group) =>  | 
|
1736  | 
let val (pc,args) = USyntax.strip_comb p  | 
|
1737  | 
in if (#1(dest_Const pc) = c)  | 
|
1738  | 
then ((args@rst, rhs)::in_group, not_in_group)  | 
|
1739  | 
else (in_group, row::not_in_group)  | 
|
1740  | 
end) rows ([],[])  | 
|
1741  | 
val col_types = Utils.take type_of (length L, #1(hd in_group))  | 
|
1742  | 
in  | 
|
1743  | 
          part{constrs = crst, rows = not_in_group,
 | 
|
1744  | 
               A = {constructor = c,
 | 
|
1745  | 
new_formals = map gv col_types,  | 
|
1746  | 
group = in_group}::A}  | 
|
1747  | 
end  | 
|
1748  | 
  in part{constrs = constructors, rows = rows, A = []}
 | 
|
1749  | 
end;  | 
|
1750  | 
||
1751  | 
||
1752  | 
||
1753  | 
(*---------------------------------------------------------------------------  | 
|
1754  | 
* Each pattern carries with it a tag (i,b) where  | 
|
1755  | 
* i is the clause it came from and  | 
|
1756  | 
* b=true indicates that clause was given by the user  | 
|
1757  | 
* (or is an instantiation of a user supplied pattern)  | 
|
1758  | 
* b=false --> i = ~1  | 
|
1759  | 
*---------------------------------------------------------------------------*)  | 
|
1760  | 
||
1761  | 
type pattern = term * (int * bool)  | 
|
1762  | 
||
1763  | 
fun pattern_map f (tm,x) = (f tm, x);  | 
|
1764  | 
||
1765  | 
fun pattern_subst theta = pattern_map (subst_free theta);  | 
|
1766  | 
||
1767  | 
val pat_of = fst;  | 
|
1768  | 
fun row_of_pat x = fst (snd x);  | 
|
1769  | 
fun given x = snd (snd x);  | 
|
1770  | 
||
1771  | 
(*---------------------------------------------------------------------------  | 
|
1772  | 
* Produce an instance of a constructor, plus genvars for its arguments.  | 
|
1773  | 
*---------------------------------------------------------------------------*)  | 
|
1774  | 
fun fresh_constr ty_match colty gv c =  | 
|
1775  | 
let val (_,Ty) = dest_Const c  | 
|
1776  | 
val L = binder_types Ty  | 
|
1777  | 
and ty = body_type Ty  | 
|
1778  | 
val ty_theta = ty_match ty colty  | 
|
1779  | 
val c' = USyntax.inst ty_theta c  | 
|
1780  | 
val gvars = map (USyntax.inst ty_theta o gv) L  | 
|
1781  | 
in (c', gvars)  | 
|
1782  | 
end;  | 
|
1783  | 
||
1784  | 
||
1785  | 
(*---------------------------------------------------------------------------  | 
|
1786  | 
* Goes through a list of rows and picks out the ones beginning with a  | 
|
1787  | 
* pattern with constructor = name.  | 
|
1788  | 
*---------------------------------------------------------------------------*)  | 
|
1789  | 
fun mk_group name rows =  | 
|
1790  | 
fold_rev (fn (row as ((prfx, p::rst), rhs)) =>  | 
|
1791  | 
fn (in_group,not_in_group) =>  | 
|
1792  | 
let val (pc,args) = USyntax.strip_comb p  | 
|
1793  | 
in if ((#1 (Term.dest_Const pc) = name) handle TERM _ => false)  | 
|
1794  | 
then (((prfx,args@rst), rhs)::in_group, not_in_group)  | 
|
1795  | 
else (in_group, row::not_in_group) end)  | 
|
1796  | 
rows ([],[]);  | 
|
1797  | 
||
1798  | 
(*---------------------------------------------------------------------------  | 
|
1799  | 
* Partition the rows. Not efficient: we should use hashing.  | 
|
1800  | 
*---------------------------------------------------------------------------*)  | 
|
1801  | 
fun partition _ _ (_,_,_,[]) = raise TFL_ERR "partition" "no rows"  | 
|
1802  | 
| partition gv ty_match  | 
|
1803  | 
(constructors, colty, res_ty, rows as (((prfx,_),_)::_)) =  | 
|
1804  | 
let val fresh = fresh_constr ty_match colty gv  | 
|
1805  | 
     fun part {constrs = [],      rows, A} = rev A
 | 
|
1806  | 
       | part {constrs = c::crst, rows, A} =
 | 
|
1807  | 
let val (c',gvars) = fresh c  | 
|
1808  | 
val (in_group, not_in_group) = mk_group (#1 (dest_Const c')) rows  | 
|
1809  | 
val in_group' =  | 
|
1810  | 
if (null in_group) (* Constructor not given *)  | 
|
1811  | 
then [((prfx, #2(fresh c)), (USyntax.ARB res_ty, (~1,false)))]  | 
|
1812  | 
else in_group  | 
|
1813  | 
in  | 
|
1814  | 
         part{constrs = crst,
 | 
|
1815  | 
rows = not_in_group,  | 
|
1816  | 
              A = {constructor = c',
 | 
|
1817  | 
new_formals = gvars,  | 
|
1818  | 
group = in_group'}::A}  | 
|
1819  | 
end  | 
|
1820  | 
in part{constrs=constructors, rows=rows, A=[]}
 | 
|
1821  | 
end;  | 
|
1822  | 
||
1823  | 
(*---------------------------------------------------------------------------  | 
|
1824  | 
* Misc. routines used in mk_case  | 
|
1825  | 
*---------------------------------------------------------------------------*)  | 
|
1826  | 
||
1827  | 
fun mk_pat (c,l) =  | 
|
1828  | 
let val L = length (binder_types (type_of c))  | 
|
1829  | 
fun build (prfx,tag,plist) =  | 
|
1830  | 
let val (args, plist') = chop L plist  | 
|
1831  | 
in (prfx,tag,list_comb(c,args)::plist') end  | 
|
1832  | 
in map build l end;  | 
|
1833  | 
||
1834  | 
fun v_to_prfx (prfx, v::pats) = (v::prfx,pats)  | 
|
1835  | 
| v_to_prfx _ = raise TFL_ERR "mk_case" "v_to_prfx";  | 
|
1836  | 
||
1837  | 
fun v_to_pats (v::prfx,tag, pats) = (prfx, tag, v::pats)  | 
|
1838  | 
| v_to_pats _ = raise TFL_ERR "mk_case" "v_to_pats";  | 
|
1839  | 
||
1840  | 
||
1841  | 
(*----------------------------------------------------------------------------  | 
|
1842  | 
* Translation of pattern terms into nested case expressions.  | 
|
1843  | 
*  | 
|
1844  | 
* This performs the translation and also builds the full set of patterns.  | 
|
1845  | 
* Thus it supports the construction of induction theorems even when an  | 
|
1846  | 
* incomplete set of patterns is given.  | 
|
1847  | 
*---------------------------------------------------------------------------*)  | 
|
1848  | 
||
1849  | 
fun mk_case ty_info ty_match usednames range_ty =  | 
|
1850  | 
let  | 
|
1851  | 
fun mk_case_fail s = raise TFL_ERR "mk_case" s  | 
|
1852  | 
val fresh_var = gvvariant usednames  | 
|
1853  | 
val divide = partition fresh_var ty_match  | 
|
| 60521 | 1854  | 
fun expand _ ty ((_,[]), _) = mk_case_fail"expand_var_row"  | 
| 60520 | 1855  | 
| expand constructors ty (row as ((prfx, p::rst), rhs)) =  | 
1856  | 
if (is_Free p)  | 
|
1857  | 
then let val fresh = fresh_constr ty_match ty fresh_var  | 
|
1858  | 
fun expnd (c,gvs) =  | 
|
1859  | 
let val capp = list_comb(c,gvs)  | 
|
1860  | 
in ((prfx, capp::rst), pattern_subst[(p,capp)] rhs)  | 
|
1861  | 
end  | 
|
1862  | 
in map expnd (map fresh constructors) end  | 
|
1863  | 
else [row]  | 
|
1864  | 
 fun mk{rows=[],...} = mk_case_fail"no rows"
 | 
|
1865  | 
   | mk{path=[], rows = ((prfx, []), (tm,tag))::_} =  (* Done *)
 | 
|
1866  | 
([(prfx,tag,[])], tm)  | 
|
1867  | 
   | mk{path=[], rows = _::_} = mk_case_fail"blunder"
 | 
|
1868  | 
   | mk{path as u::rstp, rows as ((prfx, []), rhs)::rst} =
 | 
|
1869  | 
        mk{path = path,
 | 
|
1870  | 
rows = ((prfx, [fresh_var(type_of u)]), rhs)::rst}  | 
|
1871  | 
   | mk{path = u::rstp, rows as ((_, p::_), _)::_} =
 | 
|
1872  | 
let val (pat_rectangle,rights) = ListPair.unzip rows  | 
|
1873  | 
val col0 = map(hd o #2) pat_rectangle  | 
|
1874  | 
in  | 
|
1875  | 
if (forall is_Free col0)  | 
|
1876  | 
then let val rights' = map (fn(v,e) => pattern_subst[(v,u)] e)  | 
|
1877  | 
(ListPair.zip (col0, rights))  | 
|
1878  | 
val pat_rectangle' = map v_to_prfx pat_rectangle  | 
|
1879  | 
              val (pref_patl,tm) = mk{path = rstp,
 | 
|
1880  | 
rows = ListPair.zip (pat_rectangle',  | 
|
1881  | 
rights')}  | 
|
1882  | 
in (map v_to_pats pref_patl, tm)  | 
|
1883  | 
end  | 
|
1884  | 
else  | 
|
1885  | 
let val pty as Type (ty_name,_) = type_of p  | 
|
1886  | 
in  | 
|
1887  | 
case (ty_info ty_name)  | 
|
1888  | 
     of NONE => mk_case_fail("Not a known datatype: "^ty_name)
 | 
|
1889  | 
      | SOME{case_const,constructors} =>
 | 
|
1890  | 
let  | 
|
1891  | 
val case_const_name = #1(dest_Const case_const)  | 
|
1892  | 
val nrows = maps (expand constructors pty) rows  | 
|
1893  | 
val subproblems = divide(constructors, pty, range_ty, nrows)  | 
|
1894  | 
val groups = map #group subproblems  | 
|
1895  | 
and new_formals = map #new_formals subproblems  | 
|
1896  | 
and constructors' = map #constructor subproblems  | 
|
1897  | 
            val news = map (fn (nf,rows) => {path = nf@rstp, rows=rows})
 | 
|
1898  | 
(ListPair.zip (new_formals, groups))  | 
|
1899  | 
val rec_calls = map mk news  | 
|
1900  | 
val (pat_rect,dtrees) = ListPair.unzip rec_calls  | 
|
1901  | 
val case_functions = map USyntax.list_mk_abs  | 
|
1902  | 
(ListPair.zip (new_formals, dtrees))  | 
|
1903  | 
val types = map type_of (case_functions@[u]) @ [range_ty]  | 
|
1904  | 
val case_const' = Const(case_const_name, list_mk_type types)  | 
|
1905  | 
val tree = list_comb(case_const', case_functions@[u])  | 
|
1906  | 
val pat_rect1 = flat (ListPair.map mk_pat (constructors', pat_rect))  | 
|
1907  | 
in (pat_rect1,tree)  | 
|
1908  | 
end  | 
|
1909  | 
end end  | 
|
1910  | 
in mk  | 
|
1911  | 
end;  | 
|
1912  | 
||
1913  | 
||
1914  | 
(* Repeated variable occurrences in a pattern are not allowed. *)  | 
|
1915  | 
fun FV_multiset tm =  | 
|
1916  | 
case (USyntax.dest_term tm)  | 
|
1917  | 
     of USyntax.VAR{Name = c, Ty = T} => [Free(c, T)]
 | 
|
1918  | 
| USyntax.CONST _ => []  | 
|
1919  | 
      | USyntax.COMB{Rator, Rand} => FV_multiset Rator @ FV_multiset Rand
 | 
|
1920  | 
| USyntax.LAMB _ => raise TFL_ERR "FV_multiset" "lambda";  | 
|
1921  | 
||
1922  | 
fun no_repeat_vars thy pat =  | 
|
1923  | 
let fun check [] = true  | 
|
1924  | 
| check (v::rst) =  | 
|
1925  | 
if member (op aconv) rst v then  | 
|
1926  | 
raise TFL_ERR "no_repeat_vars"  | 
|
1927  | 
(quote (#1 (dest_Free v)) ^  | 
|
1928  | 
" occurs repeatedly in the pattern " ^  | 
|
1929  | 
quote (Syntax.string_of_term_global thy pat))  | 
|
1930  | 
else check rst  | 
|
1931  | 
in check (FV_multiset pat)  | 
|
1932  | 
end;  | 
|
1933  | 
||
1934  | 
fun dest_atom (Free p) = p  | 
|
1935  | 
| dest_atom (Const p) = p  | 
|
1936  | 
| dest_atom _ = raise TFL_ERR "dest_atom" "function name not an identifier";  | 
|
1937  | 
||
1938  | 
fun same_name (p,q) = #1(dest_atom p) = #1(dest_atom q);  | 
|
1939  | 
||
1940  | 
local fun mk_functional_err s = raise TFL_ERR "mk_functional" s  | 
|
1941  | 
fun single [_$_] =  | 
|
1942  | 
mk_functional_err "recdef does not allow currying"  | 
|
1943  | 
| single [f] = f  | 
|
1944  | 
| single fs =  | 
|
1945  | 
(*multiple function names?*)  | 
|
1946  | 
if length (distinct same_name fs) < length fs  | 
|
1947  | 
then mk_functional_err  | 
|
1948  | 
"The function being declared appears with multiple types"  | 
|
1949  | 
else mk_functional_err  | 
|
1950  | 
(string_of_int (length fs) ^  | 
|
1951  | 
" distinct function names being declared")  | 
|
1952  | 
in  | 
|
1953  | 
fun mk_functional thy clauses =  | 
|
1954  | 
let val (L,R) = ListPair.unzip (map HOLogic.dest_eq clauses  | 
|
1955  | 
handle TERM _ => raise TFL_ERR "mk_functional"  | 
|
1956  | 
"recursion equations must use the = relation")  | 
|
1957  | 
val (funcs,pats) = ListPair.unzip (map (fn (t$u) =>(t,u)) L)  | 
|
1958  | 
val atom = single (distinct (op aconv) funcs)  | 
|
1959  | 
val (fname,ftype) = dest_atom atom  | 
|
| 60521 | 1960  | 
val _ = map (no_repeat_vars thy) pats  | 
| 60520 | 1961  | 
val rows = ListPair.zip (map (fn x => ([]:term list,[x])) pats,  | 
1962  | 
map_index (fn (i, t) => (t,(i,true))) R)  | 
|
1963  | 
val names = List.foldr Misc_Legacy.add_term_names [] R  | 
|
1964  | 
val atype = type_of(hd pats)  | 
|
1965  | 
and aname = singleton (Name.variant_list names) "a"  | 
|
1966  | 
val a = Free(aname,atype)  | 
|
1967  | 
val ty_info = Thry.match_info thy  | 
|
1968  | 
val ty_match = Thry.match_type thy  | 
|
1969  | 
val range_ty = type_of (hd R)  | 
|
1970  | 
val (patts, case_tm) = mk_case ty_info ty_match (aname::names) range_ty  | 
|
1971  | 
                                    {path=[a], rows=rows}
 | 
|
1972  | 
val patts1 = map (fn (_,tag,[pat]) => (pat,tag)) patts  | 
|
1973  | 
handle Match => mk_functional_err "error in pattern-match translation"  | 
|
1974  | 
val patts2 = Library.sort (Library.int_ord o apply2 row_of_pat) patts1  | 
|
1975  | 
val finals = map row_of_pat patts2  | 
|
1976  | 
val originals = map (row_of_pat o #2) rows  | 
|
| 60521 | 1977  | 
val _ = case (subtract (op =) finals originals)  | 
| 60520 | 1978  | 
of [] => ()  | 
1979  | 
| L => mk_functional_err  | 
|
1980  | 
 ("The following clauses are redundant (covered by preceding clauses): " ^
 | 
|
1981  | 
commas (map (fn i => string_of_int (i + 1)) L))  | 
|
1982  | 
 in {functional = Abs(Long_Name.base_name fname, ftype,
 | 
|
1983  | 
abstract_over (atom, absfree (aname,atype) case_tm)),  | 
|
1984  | 
pats = patts2}  | 
|
1985  | 
end end;  | 
|
1986  | 
||
1987  | 
||
1988  | 
(*----------------------------------------------------------------------------  | 
|
1989  | 
*  | 
|
1990  | 
* PRINCIPLES OF DEFINITION  | 
|
1991  | 
*  | 
|
1992  | 
*---------------------------------------------------------------------------*)  | 
|
1993  | 
||
1994  | 
||
1995  | 
(*For Isabelle, the lhs of a definition must be a constant.*)  | 
|
1996  | 
fun const_def sign (c, Ty, rhs) =  | 
|
1997  | 
singleton (Syntax.check_terms (Proof_Context.init_global sign))  | 
|
| 69593 | 1998  | 
(Const(\<^const_name>\<open>Pure.eq\<close>,dummyT) $ Const(c,Ty) $ rhs);  | 
| 60520 | 1999  | 
|
2000  | 
(*Make all TVars available for instantiation by adding a ? to the front*)  | 
|
2001  | 
fun poly_tvars (Type(a,Ts)) = Type(a, map (poly_tvars) Ts)  | 
|
2002  | 
  | poly_tvars (TFree (a,sort)) = TVar (("?" ^ a, 0), sort)
 | 
|
2003  | 
  | poly_tvars (TVar ((a,i),sort)) = TVar (("?" ^ a, i+1), sort);
 | 
|
2004  | 
||
2005  | 
local  | 
|
2006  | 
val f_eq_wfrec_R_M =  | 
|
| 60522 | 2007  | 
    #ant(USyntax.dest_imp(#2(USyntax.strip_forall (concl @{thm tfl_wfrec}))))
 | 
| 60520 | 2008  | 
  val {lhs=f, rhs} = USyntax.dest_eq f_eq_wfrec_R_M
 | 
| 60521 | 2009  | 
val _ = dest_Free f  | 
| 60520 | 2010  | 
val (wfrec,_) = USyntax.strip_comb rhs  | 
2011  | 
in  | 
|
2012  | 
||
2013  | 
fun wfrec_definition0 fid R (functional as Abs(x, Ty, _)) thy =  | 
|
2014  | 
let  | 
|
2015  | 
val def_name = Thm.def_name (Long_Name.base_name fid)  | 
|
2016  | 
val wfrec_R_M = map_types poly_tvars (wfrec $ map_types poly_tvars R) $ functional  | 
|
2017  | 
val def_term = const_def thy (fid, Ty, wfrec_R_M)  | 
|
2018  | 
val ([def], thy') =  | 
|
2019  | 
Global_Theory.add_defs false [Thm.no_attributes (Binding.name def_name, def_term)] thy  | 
|
2020  | 
in (def, thy') end;  | 
|
2021  | 
||
2022  | 
end;  | 
|
2023  | 
||
2024  | 
||
2025  | 
||
2026  | 
(*---------------------------------------------------------------------------  | 
|
2027  | 
* This structure keeps track of congruence rules that aren't derived  | 
|
2028  | 
* from a datatype definition.  | 
|
2029  | 
*---------------------------------------------------------------------------*)  | 
|
2030  | 
fun extraction_thms thy =  | 
|
2031  | 
 let val {case_rewrites,case_congs} = Thry.extract_info thy
 | 
|
2032  | 
in (case_rewrites, case_congs)  | 
|
2033  | 
end;  | 
|
2034  | 
||
2035  | 
||
2036  | 
(*---------------------------------------------------------------------------  | 
|
2037  | 
* Pair patterns with termination conditions. The full list of patterns for  | 
|
2038  | 
* a definition is merged with the TCs arising from the user-given clauses.  | 
|
2039  | 
* There can be fewer clauses than the full list, if the user omitted some  | 
|
2040  | 
* cases. This routine is used to prepare input for mk_induction.  | 
|
2041  | 
*---------------------------------------------------------------------------*)  | 
|
2042  | 
fun merge full_pats TCs =  | 
|
2043  | 
let fun insert (p,TCs) =  | 
|
2044  | 
let fun insrt ((x as (h,[]))::rst) =  | 
|
2045  | 
if (p aconv h) then (p,TCs)::rst else x::insrt rst  | 
|
2046  | 
| insrt (x::rst) = x::insrt rst  | 
|
2047  | 
| insrt[] = raise TFL_ERR "merge.insert" "pattern not found"  | 
|
2048  | 
in insrt end  | 
|
2049  | 
fun pass ([],ptcl_final) = ptcl_final  | 
|
2050  | 
| pass (ptcs::tcl, ptcl) = pass(tcl, insert ptcs ptcl)  | 
|
2051  | 
in  | 
|
2052  | 
pass (TCs, map (fn p => (p,[])) full_pats)  | 
|
2053  | 
end;  | 
|
2054  | 
||
2055  | 
||
2056  | 
fun givens pats = map pat_of (filter given pats);  | 
|
2057  | 
||
2058  | 
fun post_definition ctxt meta_tflCongs (def, pats) =  | 
|
2059  | 
let val thy = Proof_Context.theory_of ctxt  | 
|
2060  | 
val tych = Thry.typecheck thy  | 
|
2061  | 
val f = #lhs(USyntax.dest_eq(concl def))  | 
|
| 60522 | 2062  | 
     val corollary = Rules.MATCH_MP @{thm tfl_wfrec} def
 | 
| 60520 | 2063  | 
val pats' = filter given pats  | 
2064  | 
val given_pats = map pat_of pats'  | 
|
2065  | 
val rows = map row_of_pat pats'  | 
|
2066  | 
val WFR = #ant(USyntax.dest_imp(concl corollary))  | 
|
2067  | 
val R = #Rand(USyntax.dest_comb WFR)  | 
|
2068  | 
val corollary' = Rules.UNDISCH corollary (* put WF R on assums *)  | 
|
2069  | 
val corollaries = map (fn pat => Rules.SPEC (tych pat) corollary') given_pats  | 
|
2070  | 
val (case_rewrites,context_congs) = extraction_thms thy  | 
|
2071  | 
(*case_ss causes minimal simplification: bodies of case expressions are  | 
|
2072  | 
not simplified. Otherwise large examples (Red-Black trees) are too  | 
|
2073  | 
slow.*)  | 
|
2074  | 
val case_simpset =  | 
|
2075  | 
put_simpset HOL_basic_ss ctxt  | 
|
2076  | 
addsimps case_rewrites  | 
|
2077  | 
|> fold (Simplifier.add_cong o #case_cong_weak o snd)  | 
|
2078  | 
(Symtab.dest (BNF_LFP_Compat.get_all thy [BNF_LFP_Compat.Keep_Nesting]))  | 
|
2079  | 
val corollaries' = map (Simplifier.simplify case_simpset) corollaries  | 
|
2080  | 
val extract =  | 
|
2081  | 
      Rules.CONTEXT_REWRITE_RULE ctxt (f, [R], @{thm cut_apply}, meta_tflCongs @ context_congs)
 | 
|
2082  | 
val (rules, TCs) = ListPair.unzip (map extract corollaries')  | 
|
| 60522 | 2083  | 
     val rules0 = map (rewrite_rule ctxt @{thms tfl_cut_def}) rules
 | 
| 60520 | 2084  | 
val mk_cond_rule = Rules.FILTER_DISCH_ALL(not o curry (op aconv) WFR)  | 
2085  | 
val rules1 = Rules.LIST_CONJ(map mk_cond_rule rules0)  | 
|
2086  | 
in  | 
|
2087  | 
 {rules = rules1,
 | 
|
2088  | 
rows = rows,  | 
|
2089  | 
full_pats_TCs = merge (map pat_of pats) (ListPair.zip (given_pats, TCs)),  | 
|
2090  | 
TCs = TCs}  | 
|
2091  | 
end;  | 
|
2092  | 
||
2093  | 
||
2094  | 
(*----------------------------------------------------------------------------  | 
|
2095  | 
*  | 
|
2096  | 
* INDUCTION THEOREM  | 
|
2097  | 
*  | 
|
2098  | 
*---------------------------------------------------------------------------*)  | 
|
2099  | 
||
2100  | 
||
2101  | 
(*------------------------ Miscellaneous function --------------------------  | 
|
2102  | 
*  | 
|
2103  | 
* [x_1,...,x_n] ?v_1...v_n. M[v_1,...,v_n]  | 
|
2104  | 
* -----------------------------------------------------------  | 
|
2105  | 
* ( M[x_1,...,x_n], [(x_i,?v_1...v_n. M[v_1,...,v_n]),  | 
|
2106  | 
* ...  | 
|
2107  | 
* (x_j,?v_n. M[x_1,...,x_(n-1),v_n])] )  | 
|
2108  | 
*  | 
|
2109  | 
* This function is totally ad hoc. Used in the production of the induction  | 
|
2110  | 
* theorem. The nchotomy theorem can have clauses that look like  | 
|
2111  | 
*  | 
|
2112  | 
* ?v1..vn. z = C vn..v1  | 
|
2113  | 
*  | 
|
2114  | 
* in which the order of quantification is not the order of occurrence of the  | 
|
2115  | 
* quantified variables as arguments to C. Since we have no control over this  | 
|
2116  | 
* aspect of the nchotomy theorem, we make the correspondence explicit by  | 
|
2117  | 
* pairing the incoming new variable with the term it gets beta-reduced into.  | 
|
2118  | 
*---------------------------------------------------------------------------*)  | 
|
2119  | 
||
2120  | 
fun alpha_ex_unroll (xlist, tm) =  | 
|
2121  | 
let val (qvars,body) = USyntax.strip_exists tm  | 
|
2122  | 
val vlist = #2 (USyntax.strip_comb (USyntax.rhs body))  | 
|
2123  | 
val plist = ListPair.zip (vlist, xlist)  | 
|
2124  | 
val args = map (the o AList.lookup (op aconv) plist) qvars  | 
|
2125  | 
handle Option.Option => raise Fail "TFL.alpha_ex_unroll: no correspondence"  | 
|
2126  | 
fun build ex [] = []  | 
|
2127  | 
| build (_$rex) (v::rst) =  | 
|
2128  | 
let val ex1 = Term.betapply(rex, v)  | 
|
2129  | 
in ex1 :: build ex1 rst  | 
|
2130  | 
end  | 
|
2131  | 
val (nex::exl) = rev (tm::build tm args)  | 
|
2132  | 
in  | 
|
2133  | 
(nex, ListPair.zip (args, rev exl))  | 
|
2134  | 
end;  | 
|
2135  | 
||
2136  | 
||
2137  | 
||
2138  | 
(*----------------------------------------------------------------------------  | 
|
2139  | 
*  | 
|
2140  | 
* PROVING COMPLETENESS OF PATTERNS  | 
|
2141  | 
*  | 
|
2142  | 
*---------------------------------------------------------------------------*)  | 
|
2143  | 
||
| 60699 | 2144  | 
fun mk_case ctxt ty_info usednames =  | 
| 60520 | 2145  | 
let  | 
| 60699 | 2146  | 
val thy = Proof_Context.theory_of ctxt  | 
| 60520 | 2147  | 
val divide = ipartition (gvvariant usednames)  | 
2148  | 
val tych = Thry.typecheck thy  | 
|
2149  | 
fun tych_binding(x,y) = (tych x, tych y)  | 
|
2150  | 
fun fail s = raise TFL_ERR "mk_case" s  | 
|
2151  | 
 fun mk{rows=[],...} = fail"no rows"
 | 
|
2152  | 
   | mk{path=[], rows = [([], (thm, bindings))]} =
 | 
|
2153  | 
Rules.IT_EXISTS ctxt (map tych_binding bindings) thm  | 
|
2154  | 
   | mk{path = u::rstp, rows as (p::_, _)::_} =
 | 
|
2155  | 
let val (pat_rectangle,rights) = ListPair.unzip rows  | 
|
2156  | 
val col0 = map hd pat_rectangle  | 
|
2157  | 
val pat_rectangle' = map tl pat_rectangle  | 
|
2158  | 
in  | 
|
2159  | 
if (forall is_Free col0) (* column 0 is all variables *)  | 
|
2160  | 
then let val rights' = map (fn ((thm,theta),v) => (thm,theta@[(u,v)]))  | 
|
2161  | 
(ListPair.zip (rights, col0))  | 
|
2162  | 
          in mk{path = rstp, rows = ListPair.zip (pat_rectangle', rights')}
 | 
|
2163  | 
end  | 
|
2164  | 
else (* column 0 is all constructors *)  | 
|
2165  | 
let val Type (ty_name,_) = type_of p  | 
|
2166  | 
in  | 
|
2167  | 
case (ty_info ty_name)  | 
|
2168  | 
     of NONE => fail("Not a known datatype: "^ty_name)
 | 
|
2169  | 
      | SOME{constructors,nchotomy} =>
 | 
|
2170  | 
let val thm' = Rules.ISPEC (tych u) nchotomy  | 
|
2171  | 
val disjuncts = USyntax.strip_disj (concl thm')  | 
|
2172  | 
val subproblems = divide(constructors, rows)  | 
|
2173  | 
val groups = map #group subproblems  | 
|
2174  | 
and new_formals = map #new_formals subproblems  | 
|
2175  | 
val existentials = ListPair.map alpha_ex_unroll  | 
|
2176  | 
(new_formals, disjuncts)  | 
|
2177  | 
val constraints = map #1 existentials  | 
|
2178  | 
val vexl = map #2 existentials  | 
|
2179  | 
fun expnd tm (pats,(th,b)) = (pats, (Rules.SUBS ctxt [Rules.ASSUME (tych tm)] th, b))  | 
|
2180  | 
            val news = map (fn (nf,rows,c) => {path = nf@rstp,
 | 
|
2181  | 
rows = map (expnd c) rows})  | 
|
2182  | 
(Utils.zip3 new_formals groups constraints)  | 
|
2183  | 
val recursive_thms = map mk news  | 
|
2184  | 
val build_exists = Library.foldr  | 
|
2185  | 
(fn((x,t), th) =>  | 
|
2186  | 
Rules.CHOOSE ctxt (tych x, Rules.ASSUME (tych t)) th)  | 
|
2187  | 
val thms' = ListPair.map build_exists (vexl, recursive_thms)  | 
|
2188  | 
val same_concls = Rules.EVEN_ORS thms'  | 
|
2189  | 
in Rules.DISJ_CASESL thm' same_concls  | 
|
2190  | 
end  | 
|
2191  | 
end end  | 
|
2192  | 
in mk  | 
|
2193  | 
end;  | 
|
2194  | 
||
2195  | 
||
| 60699 | 2196  | 
fun complete_cases ctxt =  | 
2197  | 
let val thy = Proof_Context.theory_of ctxt  | 
|
| 60520 | 2198  | 
val tych = Thry.typecheck thy  | 
2199  | 
val ty_info = Thry.induct_info thy  | 
|
2200  | 
in fn pats =>  | 
|
2201  | 
let val names = List.foldr Misc_Legacy.add_term_names [] pats  | 
|
2202  | 
val T = type_of (hd pats)  | 
|
2203  | 
val aname = singleton (Name.variant_list names) "a"  | 
|
2204  | 
val vname = singleton (Name.variant_list (aname::names)) "v"  | 
|
2205  | 
val a = Free (aname, T)  | 
|
2206  | 
val v = Free (vname, T)  | 
|
2207  | 
val a_eq_v = HOLogic.mk_eq(a,v)  | 
|
| 60781 | 2208  | 
     val ex_th0 = Rules.EXISTS ctxt (tych (USyntax.mk_exists{Bvar=v,Body=a_eq_v}), tych a)
 | 
| 60520 | 2209  | 
(Rules.REFL (tych a))  | 
2210  | 
val th0 = Rules.ASSUME (tych a_eq_v)  | 
|
2211  | 
val rows = map (fn x => ([x], (th0,[]))) pats  | 
|
2212  | 
in  | 
|
2213  | 
Rules.GEN ctxt (tych a)  | 
|
2214  | 
(Rules.RIGHT_ASSOC ctxt  | 
|
2215  | 
(Rules.CHOOSE ctxt (tych v, ex_th0)  | 
|
| 60699 | 2216  | 
(mk_case ctxt ty_info (vname::aname::names)  | 
2217  | 
                 {path=[v], rows=rows})))
 | 
|
| 60520 | 2218  | 
end end;  | 
2219  | 
||
2220  | 
||
2221  | 
(*---------------------------------------------------------------------------  | 
|
2222  | 
* Constructing induction hypotheses: one for each recursive call.  | 
|
2223  | 
*  | 
|
2224  | 
* Note. R will never occur as a variable in the ind_clause, because  | 
|
2225  | 
* to do so, it would have to be from a nested definition, and we don't  | 
|
2226  | 
* allow nested defns to have R variable.  | 
|
2227  | 
*  | 
|
2228  | 
* Note. When the context is empty, there can be no local variables.  | 
|
2229  | 
*---------------------------------------------------------------------------*)  | 
|
2230  | 
||
2231  | 
local infix 5 ==>  | 
|
2232  | 
      fun (tm1 ==> tm2) = USyntax.mk_imp{ant = tm1, conseq = tm2}
 | 
|
2233  | 
in  | 
|
2234  | 
fun build_ih f (P,SV) (pat,TCs) =  | 
|
2235  | 
let val pat_vars = USyntax.free_vars_lr pat  | 
|
2236  | 
val globals = pat_vars@SV  | 
|
2237  | 
fun nested tm = is_some (USyntax.find_term (curry (op aconv) f) tm)  | 
|
2238  | 
fun dest_TC tm =  | 
|
2239  | 
let val (cntxt,R_y_pat) = USyntax.strip_imp(#2(USyntax.strip_forall tm))  | 
|
2240  | 
val (R,y,_) = USyntax.dest_relation R_y_pat  | 
|
2241  | 
val P_y = if (nested tm) then R_y_pat ==> P$y else P$y  | 
|
2242  | 
in case cntxt  | 
|
2243  | 
of [] => (P_y, (tm,[]))  | 
|
2244  | 
| _ => let  | 
|
2245  | 
val imp = USyntax.list_mk_conj cntxt ==> P_y  | 
|
2246  | 
val lvs = subtract (op aconv) globals (USyntax.free_vars_lr imp)  | 
|
2247  | 
val locals = #2(Utils.pluck (curry (op aconv) P) lvs) handle Utils.ERR _ => lvs  | 
|
2248  | 
in (USyntax.list_mk_forall(locals,imp), (tm,locals)) end  | 
|
2249  | 
end  | 
|
2250  | 
in case TCs  | 
|
2251  | 
of [] => (USyntax.list_mk_forall(pat_vars, P$pat), [])  | 
|
2252  | 
| _ => let val (ihs, TCs_locals) = ListPair.unzip(map dest_TC TCs)  | 
|
2253  | 
val ind_clause = USyntax.list_mk_conj ihs ==> P$pat  | 
|
2254  | 
in (USyntax.list_mk_forall(pat_vars,ind_clause), TCs_locals)  | 
|
2255  | 
end  | 
|
2256  | 
end  | 
|
2257  | 
end;  | 
|
2258  | 
||
2259  | 
(*---------------------------------------------------------------------------  | 
|
2260  | 
* This function makes good on the promise made in "build_ih".  | 
|
2261  | 
*  | 
|
2262  | 
* Input is tm = "(!y. R y pat ==> P y) ==> P pat",  | 
|
2263  | 
* TCs = TC_1[pat] ... TC_n[pat]  | 
|
2264  | 
* thm = ih1 /\ ... /\ ih_n |- ih[pat]  | 
|
2265  | 
*---------------------------------------------------------------------------*)  | 
|
2266  | 
fun prove_case ctxt f (tm,TCs_locals,thm) =  | 
|
2267  | 
let val tych = Thry.typecheck (Proof_Context.theory_of ctxt)  | 
|
2268  | 
val antc = tych(#ant(USyntax.dest_imp tm))  | 
|
2269  | 
val thm' = Rules.SPEC_ALL thm  | 
|
2270  | 
fun nested tm = is_some (USyntax.find_term (curry (op aconv) f) tm)  | 
|
2271  | 
fun get_cntxt TC = tych(#ant(USyntax.dest_imp(#2(USyntax.strip_forall(concl TC)))))  | 
|
2272  | 
fun mk_ih ((TC,locals),th2,nested) =  | 
|
2273  | 
Rules.GENL ctxt (map tych locals)  | 
|
2274  | 
(if nested then Rules.DISCH (get_cntxt TC) th2 handle Utils.ERR _ => th2  | 
|
2275  | 
else if USyntax.is_imp (concl TC) then Rules.IMP_TRANS TC th2  | 
|
2276  | 
else Rules.MP th2 TC)  | 
|
2277  | 
in  | 
|
2278  | 
Rules.DISCH antc  | 
|
2279  | 
(if USyntax.is_imp(concl thm') (* recursive calls in this clause *)  | 
|
2280  | 
then let val th1 = Rules.ASSUME antc  | 
|
2281  | 
val TCs = map #1 TCs_locals  | 
|
2282  | 
val ylist = map (#2 o USyntax.dest_relation o #2 o USyntax.strip_imp o  | 
|
2283  | 
#2 o USyntax.strip_forall) TCs  | 
|
2284  | 
val TClist = map (fn(TC,lvs) => (Rules.SPEC_ALL(Rules.ASSUME(tych TC)),lvs))  | 
|
2285  | 
TCs_locals  | 
|
2286  | 
val th2list = map (fn t => Rules.SPEC (tych t) th1) ylist  | 
|
2287  | 
val nlist = map nested TCs  | 
|
2288  | 
val triples = Utils.zip3 TClist th2list nlist  | 
|
2289  | 
val Pylist = map mk_ih triples  | 
|
2290  | 
in Rules.MP thm' (Rules.LIST_CONJ Pylist) end  | 
|
2291  | 
else thm')  | 
|
2292  | 
end;  | 
|
2293  | 
||
2294  | 
||
2295  | 
(*---------------------------------------------------------------------------  | 
|
2296  | 
*  | 
|
2297  | 
* x = (v1,...,vn) |- M[x]  | 
|
2298  | 
* ---------------------------------------------  | 
|
2299  | 
* ?v1 ... vn. x = (v1,...,vn) |- M[x]  | 
|
2300  | 
*  | 
|
2301  | 
*---------------------------------------------------------------------------*)  | 
|
2302  | 
fun LEFT_ABS_VSTRUCT ctxt tych thm =  | 
|
2303  | 
let fun CHOOSER v (tm,thm) =  | 
|
2304  | 
        let val ex_tm = USyntax.mk_exists{Bvar=v,Body=tm}
 | 
|
2305  | 
in (ex_tm, Rules.CHOOSE ctxt (tych v, Rules.ASSUME (tych ex_tm)) thm)  | 
|
2306  | 
end  | 
|
2307  | 
val [veq] = filter (can USyntax.dest_eq) (#1 (Rules.dest_thm thm))  | 
|
2308  | 
      val {lhs,rhs} = USyntax.dest_eq veq
 | 
|
2309  | 
val L = USyntax.free_vars_lr rhs  | 
|
2310  | 
in #2 (fold_rev CHOOSER L (veq,thm)) end;  | 
|
2311  | 
||
2312  | 
||
2313  | 
(*----------------------------------------------------------------------------  | 
|
2314  | 
* Input : f, R, and [(pat1,TCs1),..., (patn,TCsn)]  | 
|
2315  | 
*  | 
|
| 60522 | 2316  | 
* Instantiates tfl_wf_induct, getting Sinduct and then tries to prove  | 
| 60520 | 2317  | 
* recursion induction (Rinduct) by proving the antecedent of Sinduct from  | 
2318  | 
* the antecedent of Rinduct.  | 
|
2319  | 
*---------------------------------------------------------------------------*)  | 
|
| 60699 | 2320  | 
fun mk_induction ctxt {fconst, R, SV, pat_TCs_list} =
 | 
2321  | 
let  | 
|
2322  | 
val thy = Proof_Context.theory_of ctxt  | 
|
| 60520 | 2323  | 
val tych = Thry.typecheck thy  | 
| 60522 | 2324  | 
    val Sinduction = Rules.UNDISCH (Rules.ISPEC (tych R) @{thm tfl_wf_induct})
 | 
| 60520 | 2325  | 
val (pats,TCsl) = ListPair.unzip pat_TCs_list  | 
| 60699 | 2326  | 
val case_thm = complete_cases ctxt pats  | 
| 60520 | 2327  | 
val domain = (type_of o hd) pats  | 
2328  | 
val Pname = singleton (Name.variant_list (List.foldr (Library.foldr Misc_Legacy.add_term_names)  | 
|
2329  | 
[] (pats::TCsl))) "P"  | 
|
2330  | 
val P = Free(Pname, domain --> HOLogic.boolT)  | 
|
2331  | 
val Sinduct = Rules.SPEC (tych P) Sinduction  | 
|
2332  | 
val Sinduct_assumf = USyntax.rand ((#ant o USyntax.dest_imp o concl) Sinduct)  | 
|
2333  | 
val Rassums_TCl' = map (build_ih fconst (P,SV)) pat_TCs_list  | 
|
2334  | 
val (Rassums,TCl') = ListPair.unzip Rassums_TCl'  | 
|
2335  | 
val Rinduct_assum = Rules.ASSUME (tych (USyntax.list_mk_conj Rassums))  | 
|
2336  | 
val cases = map (fn pat => Term.betapply (Sinduct_assumf, pat)) pats  | 
|
2337  | 
val tasks = Utils.zip3 cases TCl' (Rules.CONJUNCTS Rinduct_assum)  | 
|
2338  | 
val proved_cases = map (prove_case ctxt fconst) tasks  | 
|
2339  | 
val v =  | 
|
2340  | 
Free (singleton  | 
|
2341  | 
(Name.variant_list (List.foldr Misc_Legacy.add_term_names [] (map concl proved_cases))) "v",  | 
|
2342  | 
domain)  | 
|
2343  | 
val vtyped = tych v  | 
|
2344  | 
val substs = map (Rules.SYM o Rules.ASSUME o tych o (curry HOLogic.mk_eq v)) pats  | 
|
2345  | 
val proved_cases1 = ListPair.map (fn (th,th') => Rules.SUBS ctxt [th]th')  | 
|
2346  | 
(substs, proved_cases)  | 
|
2347  | 
val abs_cases = map (LEFT_ABS_VSTRUCT ctxt tych) proved_cases1  | 
|
2348  | 
val dant = Rules.GEN ctxt vtyped (Rules.DISJ_CASESL (Rules.ISPEC vtyped case_thm) abs_cases)  | 
|
2349  | 
val dc = Rules.MP Sinduct dant  | 
|
2350  | 
val Parg_ty = type_of(#Bvar(USyntax.dest_forall(concl dc)))  | 
|
2351  | 
val vars = map (gvvariant[Pname]) (USyntax.strip_prod_type Parg_ty)  | 
|
2352  | 
val dc' = fold_rev (Rules.GEN ctxt o tych) vars  | 
|
2353  | 
(Rules.SPEC (tych(USyntax.mk_vstruct Parg_ty vars)) dc)  | 
|
2354  | 
in  | 
|
2355  | 
Rules.GEN ctxt (tych P) (Rules.DISCH (tych(concl Rinduct_assum)) dc')  | 
|
2356  | 
end  | 
|
2357  | 
handle Utils.ERR _ => raise TFL_ERR "mk_induction" "failed derivation";  | 
|
2358  | 
||
2359  | 
||
2360  | 
||
2361  | 
(*---------------------------------------------------------------------------  | 
|
2362  | 
*  | 
|
2363  | 
* POST PROCESSING  | 
|
2364  | 
*  | 
|
2365  | 
*---------------------------------------------------------------------------*)  | 
|
2366  | 
||
2367  | 
||
2368  | 
fun simplify_induction thy hth ind =  | 
|
2369  | 
let val tych = Thry.typecheck thy  | 
|
2370  | 
val (asl,_) = Rules.dest_thm ind  | 
|
2371  | 
val (_,tc_eq_tc') = Rules.dest_thm hth  | 
|
2372  | 
val tc = USyntax.lhs tc_eq_tc'  | 
|
2373  | 
fun loop [] = ind  | 
|
2374  | 
| loop (asm::rst) =  | 
|
2375  | 
if (can (Thry.match_term thy asm) tc)  | 
|
2376  | 
then Rules.UNDISCH  | 
|
2377  | 
(Rules.MATCH_MP  | 
|
| 60522 | 2378  | 
                     (Rules.MATCH_MP @{thm tfl_simp_thm} (Rules.DISCH (tych asm) ind))
 | 
| 60520 | 2379  | 
hth)  | 
2380  | 
else loop rst  | 
|
2381  | 
in loop asl  | 
|
2382  | 
end;  | 
|
2383  | 
||
2384  | 
||
2385  | 
(*---------------------------------------------------------------------------  | 
|
2386  | 
* The termination condition is an antecedent to the rule, and an  | 
|
2387  | 
* assumption to the theorem.  | 
|
2388  | 
*---------------------------------------------------------------------------*)  | 
|
2389  | 
fun elim_tc tcthm (rule,induction) =  | 
|
2390  | 
(Rules.MP rule tcthm, Rules.PROVE_HYP tcthm induction)  | 
|
2391  | 
||
2392  | 
||
2393  | 
fun trace_thms ctxt s L =  | 
|
| 61268 | 2394  | 
if !trace then writeln (cat_lines (s :: map (Thm.string_of_thm ctxt) L))  | 
| 60520 | 2395  | 
else ();  | 
2396  | 
||
2397  | 
fun trace_cterm ctxt s ct =  | 
|
2398  | 
if !trace then  | 
|
2399  | 
writeln (cat_lines [s, Syntax.string_of_term ctxt (Thm.term_of ct)])  | 
|
2400  | 
else ();  | 
|
2401  | 
||
2402  | 
||
2403  | 
fun postprocess ctxt strict {wf_tac, terminator, simplifier} {rules,induction,TCs} =
 | 
|
2404  | 
let  | 
|
2405  | 
val thy = Proof_Context.theory_of ctxt;  | 
|
2406  | 
val tych = Thry.typecheck thy;  | 
|
2407  | 
||
2408  | 
(*---------------------------------------------------------------------  | 
|
2409  | 
* Attempt to eliminate WF condition. It's the only assumption of rules  | 
|
2410  | 
*---------------------------------------------------------------------*)  | 
|
2411  | 
val (rules1,induction1) =  | 
|
2412  | 
let val thm =  | 
|
2413  | 
Rules.prove ctxt strict (HOLogic.mk_Trueprop (hd(#1(Rules.dest_thm rules))), wf_tac)  | 
|
2414  | 
in (Rules.PROVE_HYP thm rules, Rules.PROVE_HYP thm induction)  | 
|
2415  | 
end handle Utils.ERR _ => (rules,induction);  | 
|
2416  | 
||
2417  | 
(*----------------------------------------------------------------------  | 
|
2418  | 
* The termination condition (tc) is simplified to |- tc = tc' (there  | 
|
2419  | 
* might not be a change!) and then 3 attempts are made:  | 
|
2420  | 
*  | 
|
| 60522 | 2421  | 
* 1. if |- tc = T, then eliminate it with tfl_eq_True; otherwise,  | 
| 60520 | 2422  | 
* 2. apply the terminator to tc'. If |- tc' = T then eliminate; else  | 
2423  | 
* 3. replace tc by tc' in both the rules and the induction theorem.  | 
|
2424  | 
*---------------------------------------------------------------------*)  | 
|
2425  | 
||
2426  | 
fun simplify_tc tc (r,ind) =  | 
|
2427  | 
let val tc1 = tych tc  | 
|
2428  | 
val _ = trace_cterm ctxt "TC before simplification: " tc1  | 
|
2429  | 
val tc_eq = simplifier tc1  | 
|
2430  | 
val _ = trace_thms ctxt "result: " [tc_eq]  | 
|
2431  | 
in  | 
|
| 60522 | 2432  | 
       elim_tc (Rules.MATCH_MP @{thm tfl_eq_True} tc_eq) (r,ind)
 | 
| 60520 | 2433  | 
handle Utils.ERR _ =>  | 
| 60522 | 2434  | 
        (elim_tc (Rules.MATCH_MP(Rules.MATCH_MP @{thm tfl_rev_eq_mp} tc_eq)
 | 
| 60520 | 2435  | 
(Rules.prove ctxt strict (HOLogic.mk_Trueprop(USyntax.rhs(concl tc_eq)),  | 
2436  | 
terminator)))  | 
|
2437  | 
(r,ind)  | 
|
2438  | 
handle Utils.ERR _ =>  | 
|
| 60522 | 2439  | 
          (Rules.UNDISCH(Rules.MATCH_MP (Rules.MATCH_MP @{thm tfl_simp_thm} r) tc_eq),
 | 
| 60520 | 2440  | 
simplify_induction thy tc_eq ind))  | 
2441  | 
end  | 
|
2442  | 
||
2443  | 
(*----------------------------------------------------------------------  | 
|
2444  | 
* Nested termination conditions are harder to get at, since they are  | 
|
2445  | 
* left embedded in the body of the function (and in induction  | 
|
2446  | 
* theorem hypotheses). Our "solution" is to simplify them, and try to  | 
|
2447  | 
* prove termination, but leave the application of the resulting theorem  | 
|
2448  | 
* to a higher level. So things go much as in "simplify_tc": the  | 
|
2449  | 
* termination condition (tc) is simplified to |- tc = tc' (there might  | 
|
2450  | 
* not be a change) and then 2 attempts are made:  | 
|
2451  | 
*  | 
|
2452  | 
* 1. if |- tc = T, then return |- tc; otherwise,  | 
|
2453  | 
* 2. apply the terminator to tc'. If |- tc' = T then return |- tc; else  | 
|
2454  | 
* 3. return |- tc = tc'  | 
|
2455  | 
*---------------------------------------------------------------------*)  | 
|
2456  | 
fun simplify_nested_tc tc =  | 
|
2457  | 
let val tc_eq = simplifier (tych (#2 (USyntax.strip_forall tc)))  | 
|
2458  | 
in  | 
|
2459  | 
Rules.GEN_ALL ctxt  | 
|
| 60522 | 2460  | 
       (Rules.MATCH_MP @{thm tfl_eq_True} tc_eq
 | 
| 60520 | 2461  | 
handle Utils.ERR _ =>  | 
| 60522 | 2462  | 
          (Rules.MATCH_MP(Rules.MATCH_MP @{thm tfl_rev_eq_mp} tc_eq)
 | 
| 60520 | 2463  | 
(Rules.prove ctxt strict (HOLogic.mk_Trueprop (USyntax.rhs(concl tc_eq)),  | 
2464  | 
terminator))  | 
|
2465  | 
handle Utils.ERR _ => tc_eq))  | 
|
2466  | 
end  | 
|
2467  | 
||
2468  | 
(*-------------------------------------------------------------------  | 
|
2469  | 
* Attempt to simplify the termination conditions in each rule and  | 
|
2470  | 
* in the induction theorem.  | 
|
2471  | 
*-------------------------------------------------------------------*)  | 
|
2472  | 
fun strip_imp tm = if USyntax.is_neg tm then ([],tm) else USyntax.strip_imp tm  | 
|
2473  | 
fun loop ([],extras,R,ind) = (rev R, ind, extras)  | 
|
2474  | 
| loop ((r,ftcs)::rst, nthms, R, ind) =  | 
|
2475  | 
let val tcs = #1(strip_imp (concl r))  | 
|
2476  | 
val extra_tcs = subtract (op aconv) tcs ftcs  | 
|
2477  | 
val extra_tc_thms = map simplify_nested_tc extra_tcs  | 
|
2478  | 
val (r1,ind1) = fold simplify_tc tcs (r,ind)  | 
|
2479  | 
val r2 = Rules.FILTER_DISCH_ALL(not o USyntax.is_WFR) r1  | 
|
2480  | 
in loop(rst, nthms@extra_tc_thms, r2::R, ind1)  | 
|
2481  | 
end  | 
|
2482  | 
val rules_tcs = ListPair.zip (Rules.CONJUNCTS rules1, TCs)  | 
|
2483  | 
val (rules2,ind2,extras) = loop(rules_tcs,[],[],induction1)  | 
|
2484  | 
in  | 
|
2485  | 
  {induction = ind2, rules = Rules.LIST_CONJ rules2, nested_tcs = extras}
 | 
|
2486  | 
end;  | 
|
2487  | 
||
2488  | 
end;  | 
|
2489  | 
||
2490  | 
||
| 60521 | 2491  | 
|
| 60520 | 2492  | 
(*** second part of main module (postprocessing of TFL definitions) ***)  | 
2493  | 
||
2494  | 
structure Tfl: TFL =  | 
|
2495  | 
struct  | 
|
2496  | 
||
2497  | 
(* misc *)  | 
|
2498  | 
||
2499  | 
(*---------------------------------------------------------------------------  | 
|
2500  | 
* Extract termination goals so that they can be put it into a goalstack, or  | 
|
2501  | 
* have a tactic directly applied to them.  | 
|
2502  | 
*--------------------------------------------------------------------------*)  | 
|
2503  | 
fun termination_goals rules =  | 
|
2504  | 
map (Type.legacy_freeze o HOLogic.dest_Trueprop)  | 
|
2505  | 
(fold_rev (union (op aconv) o Thm.prems_of) rules []);  | 
|
2506  | 
||
2507  | 
(*---------------------------------------------------------------------------  | 
|
2508  | 
* Three postprocessors are applied to the definition. It  | 
|
2509  | 
* attempts to prove wellfoundedness of the given relation, simplifies the  | 
|
2510  | 
* non-proved termination conditions, and finally attempts to prove the  | 
|
2511  | 
* simplified termination conditions.  | 
|
2512  | 
*--------------------------------------------------------------------------*)  | 
|
2513  | 
fun std_postprocessor ctxt strict wfs =  | 
|
2514  | 
Prim.postprocess ctxt strict  | 
|
| 60774 | 2515  | 
   {wf_tac = REPEAT (ares_tac ctxt wfs 1),
 | 
| 60520 | 2516  | 
terminator =  | 
2517  | 
asm_simp_tac ctxt 1  | 
|
2518  | 
THEN TRY (Arith_Data.arith_tac ctxt 1 ORELSE  | 
|
2519  | 
        fast_force_tac (ctxt addSDs @{thms not0_implies_Suc}) 1),
 | 
|
2520  | 
simplifier = Rules.simpl_conv ctxt []};  | 
|
2521  | 
||
2522  | 
||
2523  | 
||
2524  | 
val concl = #2 o Rules.dest_thm;  | 
|
2525  | 
||
2526  | 
(*---------------------------------------------------------------------------  | 
|
2527  | 
* Postprocess a definition made by "define". This is a separate stage of  | 
|
2528  | 
* processing from the definition stage.  | 
|
2529  | 
*---------------------------------------------------------------------------*)  | 
|
2530  | 
local  | 
|
2531  | 
||
2532  | 
(* The rest of these local definitions are for the tricky nested case *)  | 
|
2533  | 
val solved = not o can USyntax.dest_eq o #2 o USyntax.strip_forall o concl  | 
|
2534  | 
||
2535  | 
fun id_thm th =  | 
|
2536  | 
   let val {lhs,rhs} = USyntax.dest_eq (#2 (USyntax.strip_forall (#2 (Rules.dest_thm th))));
 | 
|
2537  | 
in lhs aconv rhs end  | 
|
2538  | 
handle Utils.ERR _ => false;  | 
|
2539  | 
||
2540  | 
val P_imp_P_eq_True = @{thm eqTrueI} RS eq_reflection;
 | 
|
2541  | 
fun mk_meta_eq r =  | 
|
2542  | 
(case Thm.concl_of r of  | 
|
| 69593 | 2543  | 
Const(\<^const_name>\<open>Pure.eq\<close>,_)$_$_ => r  | 
2544  | 
| _ $(Const(\<^const_name>\<open>HOL.eq\<close>,_)$_$_) => r RS eq_reflection  | 
|
| 60520 | 2545  | 
| _ => r RS P_imp_P_eq_True)  | 
2546  | 
||
2547  | 
(*Is this the best way to invoke the simplifier??*)  | 
|
2548  | 
fun rewrite ctxt L = rewrite_rule ctxt (map mk_meta_eq (filter_out id_thm L))  | 
|
2549  | 
||
2550  | 
fun join_assums ctxt th =  | 
|
2551  | 
let val tych = Thm.cterm_of ctxt  | 
|
2552  | 
      val {lhs,rhs} = USyntax.dest_eq(#2 (USyntax.strip_forall (concl th)))
 | 
|
2553  | 
val cntxtl = (#1 o USyntax.strip_imp) lhs (* cntxtl should = cntxtr *)  | 
|
2554  | 
val cntxtr = (#1 o USyntax.strip_imp) rhs (* but union is solider *)  | 
|
2555  | 
val cntxt = union (op aconv) cntxtl cntxtr  | 
|
2556  | 
in  | 
|
2557  | 
Rules.GEN_ALL ctxt  | 
|
2558  | 
(Rules.DISCH_ALL  | 
|
2559  | 
(rewrite ctxt (map (Rules.ASSUME o tych) cntxt) (Rules.SPEC_ALL th)))  | 
|
2560  | 
end  | 
|
2561  | 
val gen_all = USyntax.gen_all  | 
|
2562  | 
in  | 
|
2563  | 
fun proof_stage ctxt strict wfs {f, R, rules, full_pats_TCs, TCs} =
 | 
|
2564  | 
let  | 
|
2565  | 
val _ = writeln "Proving induction theorem ..."  | 
|
2566  | 
val ind =  | 
|
| 60699 | 2567  | 
Prim.mk_induction ctxt  | 
| 60520 | 2568  | 
        {fconst=f, R=R, SV=[], pat_TCs_list=full_pats_TCs}
 | 
2569  | 
val _ = writeln "Postprocessing ...";  | 
|
2570  | 
    val {rules, induction, nested_tcs} =
 | 
|
2571  | 
      std_postprocessor ctxt strict wfs {rules=rules, induction=ind, TCs=TCs}
 | 
|
2572  | 
in  | 
|
2573  | 
case nested_tcs  | 
|
2574  | 
  of [] => {induction=induction, rules=rules,tcs=[]}
 | 
|
| 60521 | 2575  | 
| L => let val _ = writeln "Simplifying nested TCs ..."  | 
| 60520 | 2576  | 
val (solved,simplified,stubborn) =  | 
2577  | 
fold_rev (fn th => fn (So,Si,St) =>  | 
|
2578  | 
if (id_thm th) then (So, Si, th::St) else  | 
|
2579  | 
if (solved th) then (th::So, Si, St)  | 
|
2580  | 
else (So, th::Si, St)) nested_tcs ([],[],[])  | 
|
2581  | 
val simplified' = map (join_assums ctxt) simplified  | 
|
2582  | 
val dummy = (Prim.trace_thms ctxt "solved =" solved;  | 
|
2583  | 
Prim.trace_thms ctxt "simplified' =" simplified')  | 
|
2584  | 
val rewr = full_simplify (ctxt addsimps (solved @ simplified'));  | 
|
2585  | 
val dummy = Prim.trace_thms ctxt "Simplifying the induction rule..." [induction]  | 
|
2586  | 
val induction' = rewr induction  | 
|
2587  | 
val dummy = Prim.trace_thms ctxt "Simplifying the recursion rules..." [rules]  | 
|
2588  | 
val rules' = rewr rules  | 
|
2589  | 
val _ = writeln "... Postprocessing finished";  | 
|
2590  | 
in  | 
|
2591  | 
          {induction = induction',
 | 
|
2592  | 
rules = rules',  | 
|
2593  | 
tcs = map (gen_all o USyntax.rhs o #2 o USyntax.strip_forall o concl)  | 
|
2594  | 
(simplified@stubborn)}  | 
|
2595  | 
end  | 
|
2596  | 
end;  | 
|
2597  | 
||
2598  | 
||
2599  | 
(*lcp: curry the predicate of the induction rule*)  | 
|
2600  | 
fun curry_rule ctxt rl =  | 
|
2601  | 
Split_Rule.split_rule_var ctxt (Term.head_of (HOLogic.dest_Trueprop (Thm.concl_of rl))) rl;  | 
|
2602  | 
||
2603  | 
(*lcp: put a theorem into Isabelle form, using meta-level connectives*)  | 
|
2604  | 
fun meta_outer ctxt =  | 
|
2605  | 
curry_rule ctxt o Drule.export_without_context o  | 
|
| 60752 | 2606  | 
rule_by_tactic ctxt  | 
2607  | 
(REPEAT (FIRSTGOAL (resolve_tac ctxt [allI, impI, conjI] ORELSE' eresolve_tac ctxt [conjE])));  | 
|
| 60520 | 2608  | 
|
2609  | 
(*Strip off the outer !P*)  | 
|
2610  | 
val spec'=  | 
|
| 69593 | 2611  | 
  Rule_Insts.read_instantiate \<^context> [((("x", 0), Position.none), "P::'b=>bool")] [] spec;
 | 
| 60520 | 2612  | 
|
| 60524 | 2613  | 
fun simplify_defn ctxt strict congs wfs pats def0 =  | 
| 60520 | 2614  | 
let  | 
| 60825 | 2615  | 
val thy = Proof_Context.theory_of ctxt;  | 
| 
67710
 
cc2db3239932
added HOLogic.mk_obj_eq convenience and eliminated some clones;
 
wenzelm 
parents: 
66251 
diff
changeset
 | 
2616  | 
val def = HOLogic.mk_obj_eq (Thm.unvarify_global thy def0)  | 
| 60520 | 2617  | 
    val {rules, rows, TCs, full_pats_TCs} = Prim.post_definition ctxt congs (def, pats)
 | 
2618  | 
    val {lhs=f,rhs} = USyntax.dest_eq (concl def)
 | 
|
2619  | 
val (_,[R,_]) = USyntax.strip_comb rhs  | 
|
| 60521 | 2620  | 
val _ = Prim.trace_thms ctxt "congs =" congs  | 
| 60520 | 2621  | 
(*the next step has caused simplifier looping in some cases*)  | 
2622  | 
    val {induction, rules, tcs} =
 | 
|
2623  | 
proof_stage ctxt strict wfs  | 
|
2624  | 
       {f = f, R = R, rules = rules,
 | 
|
2625  | 
full_pats_TCs = full_pats_TCs,  | 
|
2626  | 
TCs = TCs}  | 
|
2627  | 
val rules' = map (Drule.export_without_context o Object_Logic.rulify_no_asm ctxt)  | 
|
2628  | 
(Rules.CONJUNCTS rules)  | 
|
2629  | 
in  | 
|
2630  | 
    {induct = meta_outer ctxt (Object_Logic.rulify_no_asm ctxt (induction RS spec')),
 | 
|
2631  | 
rules = ListPair.zip(rules', rows),  | 
|
2632  | 
tcs = (termination_goals rules') @ tcs}  | 
|
2633  | 
end  | 
|
2634  | 
  handle Utils.ERR {mesg,func,module} =>
 | 
|
2635  | 
error (mesg ^ "\n (In TFL function " ^ module ^ "." ^ func ^ ")");  | 
|
2636  | 
||
2637  | 
||
2638  | 
(* Derive the initial equations from the case-split rules to meet the  | 
|
2639  | 
users specification of the recursive function. *)  | 
|
2640  | 
local  | 
|
2641  | 
fun get_related_thms i =  | 
|
2642  | 
map_filter ((fn (r,x) => if x = i then SOME r else NONE));  | 
|
2643  | 
||
| 60521 | 2644  | 
fun solve_eq _ (_, [], _) = error "derive_init_eqs: missing rules"  | 
2645  | 
| solve_eq _ (_, [a], i) = [(a, i)]  | 
|
| 60520 | 2646  | 
| solve_eq ctxt (th, splitths, i) =  | 
2647  | 
(writeln "Proving unsplit equation...";  | 
|
2648  | 
[((Drule.export_without_context o Object_Logic.rulify_no_asm ctxt)  | 
|
2649  | 
(CaseSplit.splitto ctxt splitths th), i)])  | 
|
2650  | 
handle ERROR s =>  | 
|
2651  | 
             (warning ("recdef (solve_eq): " ^ s); map (fn x => (x,i)) splitths);
 | 
|
2652  | 
in  | 
|
2653  | 
fun derive_init_eqs ctxt rules eqs =  | 
|
2654  | 
map (Thm.trivial o Thm.cterm_of ctxt o HOLogic.mk_Trueprop) eqs  | 
|
2655  | 
|> map_index (fn (i, e) => solve_eq ctxt (e, (get_related_thms i rules), i))  | 
|
2656  | 
|> flat;  | 
|
2657  | 
end;  | 
|
2658  | 
||
2659  | 
||
2660  | 
(*---------------------------------------------------------------------------  | 
|
2661  | 
* Defining a function with an associated termination relation.  | 
|
2662  | 
*---------------------------------------------------------------------------*)  | 
|
2663  | 
fun define_i strict congs wfs fid R eqs ctxt =  | 
|
2664  | 
let  | 
|
2665  | 
val thy = Proof_Context.theory_of ctxt  | 
|
2666  | 
    val {functional, pats} = Prim.mk_functional thy eqs
 | 
|
2667  | 
val (def, thy') = Prim.wfrec_definition0 fid R functional thy  | 
|
2668  | 
val ctxt' = Proof_Context.transfer thy' ctxt  | 
|
2669  | 
val (lhs, _) = Logic.dest_equals (Thm.prop_of def)  | 
|
| 60524 | 2670  | 
    val {induct, rules, tcs} = simplify_defn ctxt' strict congs wfs pats def
 | 
| 60520 | 2671  | 
val rules' = if strict then derive_init_eqs ctxt' rules eqs else rules  | 
2672  | 
  in ({lhs = lhs, rules = rules', induct = induct, tcs = tcs}, ctxt') end;
 | 
|
2673  | 
||
2674  | 
fun define strict congs wfs fid R seqs ctxt =  | 
|
2675  | 
define_i strict congs wfs fid  | 
|
2676  | 
(Syntax.read_term ctxt R) (map (Syntax.read_term ctxt) seqs) ctxt  | 
|
2677  | 
      handle Utils.ERR {mesg,...} => error mesg;
 | 
|
2678  | 
||
2679  | 
end;  | 
|
2680  | 
||
2681  | 
end;  | 
|
2682  | 
||
2683  | 
||
| 60521 | 2684  | 
|
| 60520 | 2685  | 
(*** wrappers for Isar ***)  | 
2686  | 
||
2687  | 
(** recdef hints **)  | 
|
2688  | 
||
2689  | 
(* type hints *)  | 
|
2690  | 
||
2691  | 
type hints = {simps: thm list, congs: (string * thm) list, wfs: thm list};
 | 
|
2692  | 
||
2693  | 
fun mk_hints (simps, congs, wfs) = {simps = simps, congs = congs, wfs = wfs}: hints;
 | 
|
2694  | 
fun map_hints f ({simps, congs, wfs}: hints) = mk_hints (f (simps, congs, wfs));
 | 
|
2695  | 
||
2696  | 
fun map_simps f = map_hints (fn (simps, congs, wfs) => (f simps, congs, wfs));  | 
|
2697  | 
fun map_congs f = map_hints (fn (simps, congs, wfs) => (simps, f congs, wfs));  | 
|
2698  | 
fun map_wfs f = map_hints (fn (simps, congs, wfs) => (simps, congs, f wfs));  | 
|
2699  | 
||
2700  | 
||
2701  | 
(* congruence rules *)  | 
|
2702  | 
||
2703  | 
local  | 
|
2704  | 
||
2705  | 
val cong_head =  | 
|
2706  | 
fst o Term.dest_Const o Term.head_of o fst o Logic.dest_equals o Thm.concl_of;  | 
|
2707  | 
||
2708  | 
fun prep_cong raw_thm =  | 
|
2709  | 
let val thm = safe_mk_meta_eq raw_thm in (cong_head thm, thm) end;  | 
|
2710  | 
||
2711  | 
in  | 
|
2712  | 
||
2713  | 
fun add_cong raw_thm congs =  | 
|
2714  | 
let  | 
|
2715  | 
val (c, thm) = prep_cong raw_thm;  | 
|
2716  | 
val _ = if AList.defined (op =) congs c  | 
|
2717  | 
      then warning ("Overwriting recdef congruence rule for " ^ quote c)
 | 
|
2718  | 
else ();  | 
|
2719  | 
in AList.update (op =) (c, thm) congs end;  | 
|
2720  | 
||
2721  | 
fun del_cong raw_thm congs =  | 
|
2722  | 
let  | 
|
| 60524 | 2723  | 
val (c, _) = prep_cong raw_thm;  | 
| 60520 | 2724  | 
val _ = if AList.defined (op =) congs c  | 
2725  | 
then ()  | 
|
2726  | 
      else warning ("No recdef congruence rule for " ^ quote c);
 | 
|
2727  | 
in AList.delete (op =) c congs end;  | 
|
2728  | 
||
2729  | 
end;  | 
|
2730  | 
||
2731  | 
||
2732  | 
||
2733  | 
(** global and local recdef data **)  | 
|
2734  | 
||
2735  | 
(* theory data *)  | 
|
2736  | 
||
2737  | 
type recdef_info = {lhs: term, simps: thm list, rules: thm list list, induct: thm, tcs: term list};
 | 
|
2738  | 
||
2739  | 
structure Data = Generic_Data  | 
|
2740  | 
(  | 
|
2741  | 
type T = recdef_info Symtab.table * hints;  | 
|
2742  | 
val empty = (Symtab.empty, mk_hints ([], [], [])): T;  | 
|
2743  | 
val extend = I;  | 
|
2744  | 
fun merge  | 
|
2745  | 
   ((tab1, {simps = simps1, congs = congs1, wfs = wfs1}),
 | 
|
2746  | 
    (tab2, {simps = simps2, congs = congs2, wfs = wfs2})) : T =
 | 
|
2747  | 
(Symtab.merge (K true) (tab1, tab2),  | 
|
2748  | 
mk_hints (Thm.merge_thms (simps1, simps2),  | 
|
2749  | 
AList.merge (op =) (K true) (congs1, congs2),  | 
|
2750  | 
Thm.merge_thms (wfs1, wfs2)));  | 
|
2751  | 
);  | 
|
2752  | 
||
2753  | 
val get_recdef = Symtab.lookup o #1 o Data.get o Context.Theory;  | 
|
2754  | 
||
2755  | 
fun put_recdef name info =  | 
|
2756  | 
(Context.theory_map o Data.map o apfst) (fn tab =>  | 
|
2757  | 
Symtab.update_new (name, info) tab  | 
|
2758  | 
      handle Symtab.DUP _ => error ("Duplicate recursive function definition " ^ quote name));
 | 
|
2759  | 
||
2760  | 
val get_hints = #2 o Data.get o Context.Proof;  | 
|
2761  | 
val map_hints = Data.map o apsnd;  | 
|
2762  | 
||
2763  | 
||
2764  | 
(* attributes *)  | 
|
2765  | 
||
2766  | 
fun attrib f = Thm.declaration_attribute (map_hints o f);  | 
|
2767  | 
||
2768  | 
val simp_add = attrib (map_simps o Thm.add_thm);  | 
|
2769  | 
val simp_del = attrib (map_simps o Thm.del_thm);  | 
|
2770  | 
val cong_add = attrib (map_congs o add_cong);  | 
|
2771  | 
val cong_del = attrib (map_congs o del_cong);  | 
|
2772  | 
val wf_add = attrib (map_wfs o Thm.add_thm);  | 
|
2773  | 
val wf_del = attrib (map_wfs o Thm.del_thm);  | 
|
2774  | 
||
2775  | 
||
2776  | 
(* modifiers *)  | 
|
2777  | 
||
2778  | 
val recdef_simpN = "recdef_simp";  | 
|
2779  | 
val recdef_congN = "recdef_cong";  | 
|
2780  | 
val recdef_wfN = "recdef_wf";  | 
|
2781  | 
||
2782  | 
val recdef_modifiers =  | 
|
| 64556 | 2783  | 
[Args.$$$ recdef_simpN -- Args.colon >> K (Method.modifier simp_add \<^here>),  | 
2784  | 
Args.$$$ recdef_simpN -- Args.add -- Args.colon >> K (Method.modifier simp_add \<^here>),  | 
|
2785  | 
Args.$$$ recdef_simpN -- Args.del -- Args.colon >> K (Method.modifier simp_del \<^here>),  | 
|
2786  | 
Args.$$$ recdef_congN -- Args.colon >> K (Method.modifier cong_add \<^here>),  | 
|
2787  | 
Args.$$$ recdef_congN -- Args.add -- Args.colon >> K (Method.modifier cong_add \<^here>),  | 
|
2788  | 
Args.$$$ recdef_congN -- Args.del -- Args.colon >> K (Method.modifier cong_del \<^here>),  | 
|
2789  | 
Args.$$$ recdef_wfN -- Args.colon >> K (Method.modifier wf_add \<^here>),  | 
|
2790  | 
Args.$$$ recdef_wfN -- Args.add -- Args.colon >> K (Method.modifier wf_add \<^here>),  | 
|
2791  | 
Args.$$$ recdef_wfN -- Args.del -- Args.colon >> K (Method.modifier wf_del \<^here>)] @  | 
|
| 60520 | 2792  | 
Clasimp.clasimp_modifiers;  | 
2793  | 
||
2794  | 
||
2795  | 
||
2796  | 
(** prepare hints **)  | 
|
2797  | 
||
2798  | 
fun prepare_hints opt_src ctxt =  | 
|
2799  | 
let  | 
|
2800  | 
val ctxt' =  | 
|
2801  | 
(case opt_src of  | 
|
2802  | 
NONE => ctxt  | 
|
2803  | 
| SOME src => #2 (Token.syntax (Method.sections recdef_modifiers) src ctxt));  | 
|
2804  | 
    val {simps, congs, wfs} = get_hints ctxt';
 | 
|
2805  | 
    val ctxt'' = ctxt' addsimps simps |> Simplifier.del_cong @{thm imp_cong};
 | 
|
2806  | 
in ((rev (map snd congs), wfs), ctxt'') end;  | 
|
2807  | 
||
2808  | 
fun prepare_hints_i () ctxt =  | 
|
2809  | 
let  | 
|
2810  | 
    val {simps, congs, wfs} = get_hints ctxt;
 | 
|
2811  | 
    val ctxt' = ctxt addsimps simps |> Simplifier.del_cong @{thm imp_cong};
 | 
|
2812  | 
in ((rev (map snd congs), wfs), ctxt') end;  | 
|
2813  | 
||
2814  | 
||
2815  | 
||
2816  | 
(** add_recdef(_i) **)  | 
|
2817  | 
||
2818  | 
fun gen_add_recdef tfl_fn prep_att prep_hints not_permissive raw_name R eq_srcs hints thy =  | 
|
2819  | 
let  | 
|
2820  | 
val _ = legacy_feature "Old 'recdef' command -- use 'fun' or 'function' instead";  | 
|
2821  | 
||
2822  | 
val name = Sign.intern_const thy raw_name;  | 
|
2823  | 
val bname = Long_Name.base_name name;  | 
|
2824  | 
    val _ = writeln ("Defining recursive function " ^ quote name ^ " ...");
 | 
|
2825  | 
||
2826  | 
val ((eq_names, eqs), raw_eq_atts) = apfst split_list (split_list eq_srcs);  | 
|
2827  | 
val eq_atts = map (map (prep_att thy)) raw_eq_atts;  | 
|
2828  | 
||
2829  | 
val ((congs, wfs), ctxt) = prep_hints hints (Proof_Context.init_global thy);  | 
|
2830  | 
(*We must remove imp_cong to prevent looping when the induction rule  | 
|
2831  | 
is simplified. Many induction rules have nested implications that would  | 
|
2832  | 
give rise to looping conditional rewriting.*)  | 
|
2833  | 
    val ({lhs, rules = rules_idx, induct, tcs}, ctxt1) =
 | 
|
2834  | 
tfl_fn not_permissive congs wfs name R eqs ctxt;  | 
|
2835  | 
val rules = (map o map) fst (partition_eq (eq_snd (op = : int * int -> bool)) rules_idx);  | 
|
2836  | 
val simp_att =  | 
|
2837  | 
if null tcs then [Simplifier.simp_add,  | 
|
| 69593 | 2838  | 
Named_Theorems.add \<^named_theorems>\<open>nitpick_simp\<close>]  | 
| 60520 | 2839  | 
else [];  | 
2840  | 
val ((simps' :: rules', [induct']), thy2) =  | 
|
2841  | 
Proof_Context.theory_of ctxt1  | 
|
2842  | 
|> Sign.add_path bname  | 
|
2843  | 
|> Global_Theory.add_thmss  | 
|
2844  | 
(((Binding.name "simps", flat rules), simp_att) :: ((eq_names ~~ rules) ~~ eq_atts))  | 
|
2845  | 
||>> Global_Theory.add_thms [((Binding.name "induct", induct), [])]  | 
|
| 
71214
 
5727bcc3c47c
proper spec_rule name via naming/binding/Morphism.binding;
 
wenzelm 
parents: 
71179 
diff
changeset
 | 
2846  | 
||> Spec_Rules.add_global (Binding.name bname) Spec_Rules.equational_recdef [lhs] (flat rules)  | 
| 
66251
 
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
 
haftmann 
parents: 
64556 
diff
changeset
 | 
2847  | 
||> null tcs ? Code.declare_default_eqns_global (map (rpair true) (flat rules));  | 
| 60520 | 2848  | 
    val result = {lhs = lhs, simps = simps', rules = rules', induct = induct', tcs = tcs};
 | 
2849  | 
val thy3 =  | 
|
2850  | 
thy2  | 
|
2851  | 
|> put_recdef name result  | 
|
2852  | 
|> Sign.parent_path;  | 
|
2853  | 
in (thy3, result) end;  | 
|
2854  | 
||
2855  | 
val add_recdef = gen_add_recdef Tfl.define Attrib.attribute_cmd_global prepare_hints;  | 
|
2856  | 
fun add_recdef_i x y z w = gen_add_recdef Tfl.define_i (K I) prepare_hints_i x y z w ();  | 
|
2857  | 
||
2858  | 
||
2859  | 
||
2860  | 
(** package setup **)  | 
|
2861  | 
||
2862  | 
(* setup theory *)  | 
|
2863  | 
||
2864  | 
val _ =  | 
|
2865  | 
Theory.setup  | 
|
| 69593 | 2866  | 
(Attrib.setup \<^binding>\<open>recdef_simp\<close> (Attrib.add_del simp_add simp_del)  | 
| 60520 | 2867  | 
"declaration of recdef simp rule" #>  | 
| 69593 | 2868  | 
Attrib.setup \<^binding>\<open>recdef_cong\<close> (Attrib.add_del cong_add cong_del)  | 
| 60520 | 2869  | 
"declaration of recdef cong rule" #>  | 
| 69593 | 2870  | 
Attrib.setup \<^binding>\<open>recdef_wf\<close> (Attrib.add_del wf_add wf_del)  | 
| 60520 | 2871  | 
"declaration of recdef wf rule");  | 
2872  | 
||
2873  | 
||
2874  | 
(* outer syntax *)  | 
|
2875  | 
||
2876  | 
val hints =  | 
|
| 69593 | 2877  | 
\<^keyword>\<open>(\<close> |--  | 
2878  | 
Parse.!!! ((Parse.token \<^keyword>\<open>hints\<close> ::: Parse.args) --| \<^keyword>\<open>)\<close>);  | 
|
| 60520 | 2879  | 
|
2880  | 
val recdef_decl =  | 
|
2881  | 
Scan.optional  | 
|
| 69593 | 2882  | 
(\<^keyword>\<open>(\<close> -- Parse.!!! (\<^keyword>\<open>permissive\<close> -- \<^keyword>\<open>)\<close>) >> K false) true --  | 
| 60520 | 2883  | 
Parse.name -- Parse.term -- Scan.repeat1 (Parse_Spec.opt_thm_name ":" -- Parse.prop)  | 
2884  | 
-- Scan.option hints  | 
|
| 61466 | 2885  | 
>> (fn ((((p, f), R), eqs), src) =>  | 
2886  | 
#1 o add_recdef p f R (map (fn ((x, y), z) => ((x, z), y)) eqs) src);  | 
|
| 60520 | 2887  | 
|
2888  | 
val _ =  | 
|
| 69593 | 2889  | 
Outer_Syntax.command \<^command_keyword>\<open>recdef\<close> "define general recursive functions (obsolete TFL)"  | 
| 60520 | 2890  | 
(recdef_decl >> Toplevel.theory);  | 
2891  | 
||
2892  | 
end;  |