| author | kleing | 
| Thu, 12 Apr 2007 02:59:44 +0200 | |
| changeset 22640 | f3a6c9389e7b | 
| parent 22578 | b0eb5652f210 | 
| child 22692 | 1e057a3f087d | 
| permissions | -rw-r--r-- | 
| 10769 | 1  | 
(* Title: TFL/tfl.ML  | 
2  | 
ID: $Id$  | 
|
3  | 
Author: Konrad Slind, Cambridge University Computer Laboratory  | 
|
4  | 
Copyright 1997 University of Cambridge  | 
|
5  | 
||
6  | 
First part of main module.  | 
|
7  | 
*)  | 
|
8  | 
||
9  | 
signature PRIM =  | 
|
10  | 
sig  | 
|
11  | 
val trace: bool ref  | 
|
| 14240 | 12  | 
val trace_thms: string -> thm list -> unit  | 
13  | 
val trace_cterms: string -> cterm list -> unit  | 
|
| 10769 | 14  | 
type pattern  | 
15  | 
  val mk_functional: theory -> term list -> {functional: term, pats: pattern list}
 | 
|
16  | 
val wfrec_definition0: theory -> string -> term -> term -> theory * thm  | 
|
17  | 
val post_definition: thm list -> theory * (thm * pattern list) ->  | 
|
| 20062 | 18  | 
   {rules: thm,
 | 
| 10769 | 19  | 
rows: int list,  | 
20  | 
TCs: term list list,  | 
|
21  | 
full_pats_TCs: (term * term list) list}  | 
|
22  | 
val wfrec_eqns: theory -> xstring -> thm list -> term list ->  | 
|
23  | 
   {WFR: term,
 | 
|
24  | 
SV: term list,  | 
|
25  | 
proto_def: term,  | 
|
26  | 
extracta: (thm * term list) list,  | 
|
27  | 
pats: pattern list}  | 
|
28  | 
val lazyR_def: theory -> xstring -> thm list -> term list ->  | 
|
29  | 
   {theory: theory,
 | 
|
30  | 
rules: thm,  | 
|
31  | 
R: term,  | 
|
32  | 
SV: term list,  | 
|
33  | 
full_pats_TCs: (term * term list) list,  | 
|
34  | 
patterns : pattern list}  | 
|
35  | 
val mk_induction: theory ->  | 
|
36  | 
    {fconst: term, R: term, SV: term list, pat_TCs_list: (term * term list) list} -> thm
 | 
|
| 11632 | 37  | 
  val postprocess: bool -> {wf_tac: tactic, terminator: tactic, simplifier: cterm -> thm}
 | 
38  | 
    -> theory -> {rules: thm, induction: thm, TCs: term list list}
 | 
|
39  | 
    -> {rules: thm, induction: thm, nested_tcs: thm list}
 | 
|
| 10769 | 40  | 
end;  | 
41  | 
||
42  | 
structure Prim: PRIM =  | 
|
43  | 
struct  | 
|
44  | 
||
45  | 
val trace = ref false;  | 
|
46  | 
||
47  | 
structure R = Rules;  | 
|
48  | 
structure S = USyntax;  | 
|
49  | 
structure U = Utils;  | 
|
50  | 
||
51  | 
||
52  | 
fun TFL_ERR func mesg = U.ERR {module = "Tfl", func = func, mesg = mesg};
 | 
|
53  | 
||
54  | 
val concl = #2 o R.dest_thm;  | 
|
55  | 
val hyp = #1 o R.dest_thm;  | 
|
56  | 
||
57  | 
val list_mk_type = U.end_itlist (curry (op -->));  | 
|
58  | 
||
59  | 
fun enumerate xs = ListPair.zip(xs, 0 upto (length xs - 1));  | 
|
60  | 
||
61  | 
fun front_last [] = raise TFL_ERR "front_last" "empty list"  | 
|
62  | 
| front_last [x] = ([],x)  | 
|
63  | 
| front_last (h::t) =  | 
|
64  | 
let val (pref,x) = front_last t  | 
|
65  | 
in  | 
|
66  | 
(h::pref,x)  | 
|
67  | 
end;  | 
|
68  | 
||
69  | 
||
70  | 
(*---------------------------------------------------------------------------  | 
|
71  | 
* The next function is common to pattern-match translation and  | 
|
72  | 
* proof of completeness of cases for the induction theorem.  | 
|
73  | 
*  | 
|
74  | 
* The curried function "gvvariant" returns a function to generate distinct  | 
|
75  | 
* variables that are guaranteed not to be in names. The names of  | 
|
76  | 
* the variables go u, v, ..., z, aa, ..., az, ... The returned  | 
|
77  | 
* function contains embedded refs!  | 
|
78  | 
*---------------------------------------------------------------------------*)  | 
|
79  | 
fun gvvariant names =  | 
|
80  | 
let val slist = ref names  | 
|
81  | 
val vname = ref "u"  | 
|
82  | 
fun new() =  | 
|
83  | 
if !vname mem_string (!slist)  | 
|
| 12902 | 84  | 
then (vname := Symbol.bump_string (!vname); new())  | 
| 10769 | 85  | 
else (slist := !vname :: !slist; !vname)  | 
86  | 
in  | 
|
87  | 
fn ty => Free(new(), ty)  | 
|
88  | 
end;  | 
|
89  | 
||
90  | 
||
91  | 
(*---------------------------------------------------------------------------  | 
|
92  | 
* Used in induction theorem production. This is the simple case of  | 
|
93  | 
* partitioning up pattern rows by the leading constructor.  | 
|
94  | 
*---------------------------------------------------------------------------*)  | 
|
95  | 
fun ipartition gv (constructors,rows) =  | 
|
96  | 
let fun pfail s = raise TFL_ERR "partition.part" s  | 
|
97  | 
      fun part {constrs = [],   rows = [],   A} = rev A
 | 
|
98  | 
        | part {constrs = [],   rows = _::_, A} = pfail"extra cases in defn"
 | 
|
99  | 
        | part {constrs = _::_, rows = [],   A} = pfail"cases missing in defn"
 | 
|
100  | 
        | part {constrs = c::crst, rows,     A} =
 | 
|
| 16505 | 101  | 
let val (c, T) = dest_Const c  | 
102  | 
val L = binder_types T  | 
|
| 10769 | 103  | 
val (in_group, not_in_group) =  | 
| 16853 | 104  | 
fold_rev (fn (row as (p::rst, rhs)) =>  | 
| 10769 | 105  | 
fn (in_group,not_in_group) =>  | 
106  | 
let val (pc,args) = S.strip_comb p  | 
|
| 16505 | 107  | 
in if (#1(dest_Const pc) = c)  | 
| 10769 | 108  | 
then ((args@rst, rhs)::in_group, not_in_group)  | 
109  | 
else (in_group, row::not_in_group)  | 
|
110  | 
end) rows ([],[])  | 
|
111  | 
val col_types = U.take type_of (length L, #1(hd in_group))  | 
|
112  | 
in  | 
|
113  | 
          part{constrs = crst, rows = not_in_group,
 | 
|
114  | 
               A = {constructor = c,
 | 
|
115  | 
new_formals = map gv col_types,  | 
|
116  | 
group = in_group}::A}  | 
|
117  | 
end  | 
|
118  | 
  in part{constrs = constructors, rows = rows, A = []}
 | 
|
119  | 
end;  | 
|
120  | 
||
121  | 
||
122  | 
||
123  | 
(*---------------------------------------------------------------------------  | 
|
124  | 
* Each pattern carries with it a tag (i,b) where  | 
|
125  | 
* i is the clause it came from and  | 
|
126  | 
* b=true indicates that clause was given by the user  | 
|
127  | 
* (or is an instantiation of a user supplied pattern)  | 
|
128  | 
* b=false --> i = ~1  | 
|
129  | 
*---------------------------------------------------------------------------*)  | 
|
130  | 
||
131  | 
type pattern = term * (int * bool)  | 
|
132  | 
||
133  | 
fun pattern_map f (tm,x) = (f tm, x);  | 
|
134  | 
||
135  | 
fun pattern_subst theta = pattern_map (subst_free theta);  | 
|
136  | 
||
137  | 
val pat_of = fst;  | 
|
138  | 
fun row_of_pat x = fst (snd x);  | 
|
139  | 
fun given x = snd (snd x);  | 
|
140  | 
||
141  | 
(*---------------------------------------------------------------------------  | 
|
142  | 
* Produce an instance of a constructor, plus genvars for its arguments.  | 
|
143  | 
*---------------------------------------------------------------------------*)  | 
|
144  | 
fun fresh_constr ty_match colty gv c =  | 
|
145  | 
let val (_,Ty) = dest_Const c  | 
|
146  | 
val L = binder_types Ty  | 
|
147  | 
and ty = body_type Ty  | 
|
148  | 
val ty_theta = ty_match ty colty  | 
|
149  | 
val c' = S.inst ty_theta c  | 
|
150  | 
val gvars = map (S.inst ty_theta o gv) L  | 
|
151  | 
in (c', gvars)  | 
|
152  | 
end;  | 
|
153  | 
||
154  | 
||
155  | 
(*---------------------------------------------------------------------------  | 
|
156  | 
* Goes through a list of rows and picks out the ones beginning with a  | 
|
| 16505 | 157  | 
* pattern with constructor = name.  | 
| 10769 | 158  | 
*---------------------------------------------------------------------------*)  | 
| 16505 | 159  | 
fun mk_group name rows =  | 
| 16853 | 160  | 
fold_rev (fn (row as ((prfx, p::rst), rhs)) =>  | 
| 10769 | 161  | 
fn (in_group,not_in_group) =>  | 
162  | 
let val (pc,args) = S.strip_comb p  | 
|
| 16505 | 163  | 
in if ((#1 (Term.dest_Const pc) = name) handle TERM _ => false)  | 
| 10769 | 164  | 
then (((prfx,args@rst), rhs)::in_group, not_in_group)  | 
165  | 
else (in_group, row::not_in_group) end)  | 
|
166  | 
rows ([],[]);  | 
|
167  | 
||
168  | 
(*---------------------------------------------------------------------------  | 
|
169  | 
* Partition the rows. Not efficient: we should use hashing.  | 
|
170  | 
*---------------------------------------------------------------------------*)  | 
|
171  | 
fun partition _ _ (_,_,_,[]) = raise TFL_ERR "partition" "no rows"  | 
|
172  | 
| partition gv ty_match  | 
|
173  | 
(constructors, colty, res_ty, rows as (((prfx,_),_)::_)) =  | 
|
174  | 
let val fresh = fresh_constr ty_match colty gv  | 
|
175  | 
     fun part {constrs = [],      rows, A} = rev A
 | 
|
176  | 
       | part {constrs = c::crst, rows, A} =
 | 
|
177  | 
let val (c',gvars) = fresh c  | 
|
| 16505 | 178  | 
val (in_group, not_in_group) = mk_group (#1 (dest_Const c')) rows  | 
| 10769 | 179  | 
val in_group' =  | 
180  | 
if (null in_group) (* Constructor not given *)  | 
|
181  | 
then [((prfx, #2(fresh c)), (S.ARB res_ty, (~1,false)))]  | 
|
182  | 
else in_group  | 
|
183  | 
in  | 
|
184  | 
         part{constrs = crst,
 | 
|
185  | 
rows = not_in_group,  | 
|
186  | 
              A = {constructor = c',
 | 
|
187  | 
new_formals = gvars,  | 
|
188  | 
group = in_group'}::A}  | 
|
189  | 
end  | 
|
190  | 
in part{constrs=constructors, rows=rows, A=[]}
 | 
|
191  | 
end;  | 
|
192  | 
||
193  | 
(*---------------------------------------------------------------------------  | 
|
194  | 
* Misc. routines used in mk_case  | 
|
195  | 
*---------------------------------------------------------------------------*)  | 
|
196  | 
||
197  | 
fun mk_pat (c,l) =  | 
|
198  | 
let val L = length (binder_types (type_of c))  | 
|
199  | 
fun build (prfx,tag,plist) =  | 
|
| 15570 | 200  | 
let val args = Library.take (L,plist)  | 
201  | 
and plist' = Library.drop(L,plist)  | 
|
| 10769 | 202  | 
in (prfx,tag,list_comb(c,args)::plist') end  | 
203  | 
in map build l end;  | 
|
204  | 
||
205  | 
fun v_to_prfx (prfx, v::pats) = (v::prfx,pats)  | 
|
206  | 
| v_to_prfx _ = raise TFL_ERR "mk_case" "v_to_prfx";  | 
|
207  | 
||
208  | 
fun v_to_pats (v::prfx,tag, pats) = (prfx, tag, v::pats)  | 
|
209  | 
| v_to_pats _ = raise TFL_ERR "mk_case" "v_to_pats";  | 
|
210  | 
||
211  | 
||
212  | 
(*----------------------------------------------------------------------------  | 
|
213  | 
* Translation of pattern terms into nested case expressions.  | 
|
214  | 
*  | 
|
215  | 
* This performs the translation and also builds the full set of patterns.  | 
|
216  | 
* Thus it supports the construction of induction theorems even when an  | 
|
217  | 
* incomplete set of patterns is given.  | 
|
218  | 
*---------------------------------------------------------------------------*)  | 
|
219  | 
||
220  | 
fun mk_case ty_info ty_match usednames range_ty =  | 
|
221  | 
let  | 
|
222  | 
fun mk_case_fail s = raise TFL_ERR "mk_case" s  | 
|
223  | 
val fresh_var = gvvariant usednames  | 
|
224  | 
val divide = partition fresh_var ty_match  | 
|
225  | 
fun expand constructors ty ((_,[]), _) = mk_case_fail"expand_var_row"  | 
|
226  | 
| expand constructors ty (row as ((prfx, p::rst), rhs)) =  | 
|
227  | 
if (is_Free p)  | 
|
228  | 
then let val fresh = fresh_constr ty_match ty fresh_var  | 
|
229  | 
fun expnd (c,gvs) =  | 
|
230  | 
let val capp = list_comb(c,gvs)  | 
|
231  | 
in ((prfx, capp::rst), pattern_subst[(p,capp)] rhs)  | 
|
232  | 
end  | 
|
233  | 
in map expnd (map fresh constructors) end  | 
|
234  | 
else [row]  | 
|
235  | 
 fun mk{rows=[],...} = mk_case_fail"no rows"
 | 
|
236  | 
   | mk{path=[], rows = ((prfx, []), (tm,tag))::_} =  (* Done *)
 | 
|
237  | 
([(prfx,tag,[])], tm)  | 
|
238  | 
   | mk{path=[], rows = _::_} = mk_case_fail"blunder"
 | 
|
239  | 
   | mk{path as u::rstp, rows as ((prfx, []), rhs)::rst} =
 | 
|
240  | 
        mk{path = path,
 | 
|
241  | 
rows = ((prfx, [fresh_var(type_of u)]), rhs)::rst}  | 
|
242  | 
   | mk{path = u::rstp, rows as ((_, p::_), _)::_} =
 | 
|
243  | 
let val (pat_rectangle,rights) = ListPair.unzip rows  | 
|
244  | 
val col0 = map(hd o #2) pat_rectangle  | 
|
245  | 
in  | 
|
246  | 
if (forall is_Free col0)  | 
|
247  | 
then let val rights' = map (fn(v,e) => pattern_subst[(v,u)] e)  | 
|
248  | 
(ListPair.zip (col0, rights))  | 
|
249  | 
val pat_rectangle' = map v_to_prfx pat_rectangle  | 
|
250  | 
              val (pref_patl,tm) = mk{path = rstp,
 | 
|
251  | 
rows = ListPair.zip (pat_rectangle',  | 
|
252  | 
rights')}  | 
|
253  | 
in (map v_to_pats pref_patl, tm)  | 
|
254  | 
end  | 
|
255  | 
else  | 
|
256  | 
let val pty as Type (ty_name,_) = type_of p  | 
|
257  | 
in  | 
|
258  | 
case (ty_info ty_name)  | 
|
| 15531 | 259  | 
     of NONE => mk_case_fail("Not a known datatype: "^ty_name)
 | 
260  | 
      | SOME{case_const,constructors} =>
 | 
|
| 10769 | 261  | 
let  | 
262  | 
val case_const_name = #1(dest_Const case_const)  | 
|
263  | 
val nrows = List.concat (map (expand constructors pty) rows)  | 
|
264  | 
val subproblems = divide(constructors, pty, range_ty, nrows)  | 
|
265  | 
val groups = map #group subproblems  | 
|
266  | 
and new_formals = map #new_formals subproblems  | 
|
267  | 
and constructors' = map #constructor subproblems  | 
|
268  | 
            val news = map (fn (nf,rows) => {path = nf@rstp, rows=rows})
 | 
|
269  | 
(ListPair.zip (new_formals, groups))  | 
|
270  | 
val rec_calls = map mk news  | 
|
271  | 
val (pat_rect,dtrees) = ListPair.unzip rec_calls  | 
|
272  | 
val case_functions = map S.list_mk_abs  | 
|
273  | 
(ListPair.zip (new_formals, dtrees))  | 
|
274  | 
val types = map type_of (case_functions@[u]) @ [range_ty]  | 
|
275  | 
val case_const' = Const(case_const_name, list_mk_type types)  | 
|
276  | 
val tree = list_comb(case_const', case_functions@[u])  | 
|
277  | 
val pat_rect1 = List.concat  | 
|
278  | 
(ListPair.map mk_pat (constructors', pat_rect))  | 
|
279  | 
in (pat_rect1,tree)  | 
|
280  | 
end  | 
|
281  | 
end end  | 
|
282  | 
in mk  | 
|
283  | 
end;  | 
|
284  | 
||
285  | 
||
286  | 
(* Repeated variable occurrences in a pattern are not allowed. *)  | 
|
287  | 
fun FV_multiset tm =  | 
|
288  | 
case (S.dest_term tm)  | 
|
| 16505 | 289  | 
     of S.VAR{Name = c, Ty = T} => [Free(c, T)]
 | 
| 10769 | 290  | 
| S.CONST _ => []  | 
291  | 
      | S.COMB{Rator, Rand} => FV_multiset Rator @ FV_multiset Rand
 | 
|
292  | 
| S.LAMB _ => raise TFL_ERR "FV_multiset" "lambda";  | 
|
293  | 
||
294  | 
fun no_repeat_vars thy pat =  | 
|
295  | 
let fun check [] = true  | 
|
296  | 
| check (v::rst) =  | 
|
| 
20088
 
bffda4cd0f79
replaced Term.variant(list) by Name.variant(_list);
 
wenzelm 
parents: 
20062 
diff
changeset
 | 
297  | 
if member (op aconv) rst v then  | 
| 10769 | 298  | 
raise TFL_ERR "no_repeat_vars"  | 
299  | 
(quote (#1 (dest_Free v)) ^  | 
|
300  | 
" occurs repeatedly in the pattern " ^  | 
|
301  | 
quote (string_of_cterm (Thry.typecheck thy pat)))  | 
|
302  | 
else check rst  | 
|
303  | 
in check (FV_multiset pat)  | 
|
304  | 
end;  | 
|
305  | 
||
306  | 
fun dest_atom (Free p) = p  | 
|
307  | 
| dest_atom (Const p) = p  | 
|
308  | 
| dest_atom _ = raise TFL_ERR "dest_atom" "function name not an identifier";  | 
|
309  | 
||
310  | 
fun same_name (p,q) = #1(dest_atom p) = #1(dest_atom q);  | 
|
311  | 
||
312  | 
local fun mk_functional_err s = raise TFL_ERR "mk_functional" s  | 
|
313  | 
fun single [_$_] =  | 
|
314  | 
mk_functional_err "recdef does not allow currying"  | 
|
315  | 
| single [f] = f  | 
|
316  | 
| single fs =  | 
|
317  | 
(*multiple function names?*)  | 
|
| 
19046
 
bc5c6c9b114e
removed distinct, renamed gen_distinct to distinct;
 
wenzelm 
parents: 
18972 
diff
changeset
 | 
318  | 
if length (distinct same_name fs) < length fs  | 
| 10769 | 319  | 
then mk_functional_err  | 
320  | 
"The function being declared appears with multiple types"  | 
|
321  | 
else mk_functional_err  | 
|
322  | 
(Int.toString (length fs) ^  | 
|
323  | 
" distinct function names being declared")  | 
|
324  | 
in  | 
|
325  | 
fun mk_functional thy clauses =  | 
|
326  | 
let val (L,R) = ListPair.unzip (map HOLogic.dest_eq clauses  | 
|
327  | 
handle TERM _ => raise TFL_ERR "mk_functional"  | 
|
328  | 
"recursion equations must use the = relation")  | 
|
329  | 
val (funcs,pats) = ListPair.unzip (map (fn (t$u) =>(t,u)) L)  | 
|
| 
19046
 
bc5c6c9b114e
removed distinct, renamed gen_distinct to distinct;
 
wenzelm 
parents: 
18972 
diff
changeset
 | 
330  | 
val atom = single (distinct (op aconv) funcs)  | 
| 10769 | 331  | 
val (fname,ftype) = dest_atom atom  | 
332  | 
val dummy = map (no_repeat_vars thy) pats  | 
|
333  | 
val rows = ListPair.zip (map (fn x => ([]:term list,[x])) pats,  | 
|
334  | 
map (fn (t,i) => (t,(i,true))) (enumerate R))  | 
|
| 
15574
 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 
skalberg 
parents: 
15570 
diff
changeset
 | 
335  | 
val names = foldr add_term_names [] R  | 
| 10769 | 336  | 
val atype = type_of(hd pats)  | 
| 
20088
 
bffda4cd0f79
replaced Term.variant(list) by Name.variant(_list);
 
wenzelm 
parents: 
20062 
diff
changeset
 | 
337  | 
and aname = Name.variant names "a"  | 
| 10769 | 338  | 
val a = Free(aname,atype)  | 
339  | 
val ty_info = Thry.match_info thy  | 
|
340  | 
val ty_match = Thry.match_type thy  | 
|
341  | 
val range_ty = type_of (hd R)  | 
|
342  | 
val (patts, case_tm) = mk_case ty_info ty_match (aname::names) range_ty  | 
|
343  | 
                                    {path=[a], rows=rows}
 | 
|
344  | 
val patts1 = map (fn (_,tag,[pat]) => (pat,tag)) patts  | 
|
345  | 
handle Match => mk_functional_err "error in pattern-match translation"  | 
|
346  | 
val patts2 = Library.sort (Library.int_ord o Library.pairself row_of_pat) patts1  | 
|
347  | 
val finals = map row_of_pat patts2  | 
|
348  | 
val originals = map (row_of_pat o #2) rows  | 
|
349  | 
val dummy = case (originals\\finals)  | 
|
350  | 
of [] => ()  | 
|
351  | 
| L => mk_functional_err  | 
|
352  | 
 ("The following clauses are redundant (covered by preceding clauses): " ^
 | 
|
353  | 
commas (map (fn i => Int.toString (i + 1)) L))  | 
|
354  | 
 in {functional = Abs(Sign.base_name fname, ftype,
 | 
|
355  | 
abstract_over (atom,  | 
|
356  | 
absfree(aname,atype, case_tm))),  | 
|
357  | 
pats = patts2}  | 
|
358  | 
end end;  | 
|
359  | 
||
360  | 
||
361  | 
(*----------------------------------------------------------------------------  | 
|
362  | 
*  | 
|
363  | 
* PRINCIPLES OF DEFINITION  | 
|
364  | 
*  | 
|
365  | 
*---------------------------------------------------------------------------*)  | 
|
366  | 
||
367  | 
||
368  | 
(*For Isabelle, the lhs of a definition must be a constant.*)  | 
|
| 16505 | 369  | 
fun mk_const_def sign (c, Ty, rhs) =  | 
| 20155 | 370  | 
Sign.infer_types (Sign.pp sign) sign (Sign.consts_of sign) (K NONE) (K NONE) Name.context false  | 
| 16505 | 371  | 
               ([Const("==",dummyT) $ Const(c,Ty) $ rhs], propT)
 | 
| 10769 | 372  | 
|> #1;  | 
373  | 
||
374  | 
(*Make all TVars available for instantiation by adding a ? to the front*)  | 
|
375  | 
fun poly_tvars (Type(a,Ts)) = Type(a, map (poly_tvars) Ts)  | 
|
376  | 
  | poly_tvars (TFree (a,sort)) = TVar (("?" ^ a, 0), sort)
 | 
|
377  | 
  | poly_tvars (TVar ((a,i),sort)) = TVar (("?" ^ a, i+1), sort);
 | 
|
378  | 
||
379  | 
local val f_eq_wfrec_R_M =  | 
|
380  | 
#ant(S.dest_imp(#2(S.strip_forall (concl Thms.WFREC_COROLLARY))))  | 
|
381  | 
      val {lhs=f, rhs} = S.dest_eq f_eq_wfrec_R_M
 | 
|
382  | 
val (fname,_) = dest_Free f  | 
|
383  | 
val (wfrec,_) = S.strip_comb rhs  | 
|
384  | 
in  | 
|
| 16505 | 385  | 
fun wfrec_definition0 thy fid R (functional as Abs(x, Ty, _)) =  | 
386  | 
let val def_name = if x<>fid then  | 
|
| 10769 | 387  | 
raise TFL_ERR "wfrec_definition0"  | 
388  | 
                                      ("Expected a definition of " ^
 | 
|
389  | 
quote fid ^ " but found one of " ^  | 
|
| 16505 | 390  | 
quote x)  | 
391  | 
else x ^ "_def"  | 
|
| 
20548
 
8ef25fe585a8
renamed Term.map_term_types to Term.map_types (cf. Term.fold_types);
 
wenzelm 
parents: 
20155 
diff
changeset
 | 
392  | 
val wfrec_R_M = map_types poly_tvars  | 
| 
 
8ef25fe585a8
renamed Term.map_term_types to Term.map_types (cf. Term.fold_types);
 
wenzelm 
parents: 
20155 
diff
changeset
 | 
393  | 
(wfrec $ map_types poly_tvars R)  | 
| 10769 | 394  | 
$ functional  | 
| 22578 | 395  | 
val def_term = mk_const_def thy (x, Ty, wfrec_R_M)  | 
| 18358 | 396  | 
val ([def], thy') = PureThy.add_defs_i false [Thm.no_attributes (def_name, def_term)] thy  | 
| 10769 | 397  | 
in (thy', def) end;  | 
398  | 
end;  | 
|
399  | 
||
400  | 
||
401  | 
||
402  | 
(*---------------------------------------------------------------------------  | 
|
403  | 
* This structure keeps track of congruence rules that aren't derived  | 
|
404  | 
* from a datatype definition.  | 
|
405  | 
*---------------------------------------------------------------------------*)  | 
|
406  | 
fun extraction_thms thy =  | 
|
407  | 
 let val {case_rewrites,case_congs} = Thry.extract_info thy
 | 
|
408  | 
in (case_rewrites, case_congs)  | 
|
409  | 
end;  | 
|
410  | 
||
411  | 
||
412  | 
(*---------------------------------------------------------------------------  | 
|
413  | 
* Pair patterns with termination conditions. The full list of patterns for  | 
|
414  | 
* a definition is merged with the TCs arising from the user-given clauses.  | 
|
415  | 
* There can be fewer clauses than the full list, if the user omitted some  | 
|
416  | 
* cases. This routine is used to prepare input for mk_induction.  | 
|
417  | 
*---------------------------------------------------------------------------*)  | 
|
418  | 
fun merge full_pats TCs =  | 
|
419  | 
let fun insert (p,TCs) =  | 
|
420  | 
let fun insrt ((x as (h,[]))::rst) =  | 
|
421  | 
if (p aconv h) then (p,TCs)::rst else x::insrt rst  | 
|
422  | 
| insrt (x::rst) = x::insrt rst  | 
|
423  | 
| insrt[] = raise TFL_ERR "merge.insert" "pattern not found"  | 
|
424  | 
in insrt end  | 
|
425  | 
fun pass ([],ptcl_final) = ptcl_final  | 
|
426  | 
| pass (ptcs::tcl, ptcl) = pass(tcl, insert ptcs ptcl)  | 
|
427  | 
in  | 
|
428  | 
pass (TCs, map (fn p => (p,[])) full_pats)  | 
|
429  | 
end;  | 
|
430  | 
||
431  | 
||
| 15570 | 432  | 
fun givens pats = map pat_of (List.filter given pats);  | 
| 10769 | 433  | 
|
434  | 
fun post_definition meta_tflCongs (theory, (def, pats)) =  | 
|
435  | 
let val tych = Thry.typecheck theory  | 
|
436  | 
val f = #lhs(S.dest_eq(concl def))  | 
|
437  | 
val corollary = R.MATCH_MP Thms.WFREC_COROLLARY def  | 
|
| 15570 | 438  | 
val pats' = List.filter given pats  | 
| 10769 | 439  | 
val given_pats = map pat_of pats'  | 
440  | 
val rows = map row_of_pat pats'  | 
|
441  | 
val WFR = #ant(S.dest_imp(concl corollary))  | 
|
442  | 
val R = #Rand(S.dest_comb WFR)  | 
|
443  | 
val corollary' = R.UNDISCH corollary (* put WF R on assums *)  | 
|
444  | 
val corollaries = map (fn pat => R.SPEC (tych pat) corollary')  | 
|
445  | 
given_pats  | 
|
446  | 
val (case_rewrites,context_congs) = extraction_thms theory  | 
|
| 14219 | 447  | 
(*case_ss causes minimal simplification: bodies of case expressions are  | 
| 16505 | 448  | 
not simplified. Otherwise large examples (Red-Black trees) are too  | 
| 14219 | 449  | 
slow.*)  | 
| 20062 | 450  | 
val case_ss = Simplifier.theory_context theory  | 
451  | 
(HOL_basic_ss addcongs  | 
|
452  | 
(map (#weak_case_cong o snd) o Symtab.dest o DatatypePackage.get_datatypes) theory addsimps case_rewrites)  | 
|
| 
14217
 
9f5679e97eac
Fixed inefficiency in post_definition by adding weak case congruence
 
berghofe 
parents: 
12902 
diff
changeset
 | 
453  | 
val corollaries' = map (Simplifier.simplify case_ss) corollaries  | 
| 10769 | 454  | 
val extract = R.CONTEXT_REWRITE_RULE  | 
455  | 
(f, [R], cut_apply, meta_tflCongs@context_congs)  | 
|
456  | 
val (rules, TCs) = ListPair.unzip (map extract corollaries')  | 
|
457  | 
val rules0 = map (rewrite_rule [Thms.CUT_DEF]) rules  | 
|
458  | 
val mk_cond_rule = R.FILTER_DISCH_ALL(not o curry (op aconv) WFR)  | 
|
459  | 
val rules1 = R.LIST_CONJ(map mk_cond_rule rules0)  | 
|
460  | 
in  | 
|
| 20062 | 461  | 
 {rules = rules1,
 | 
| 10769 | 462  | 
rows = rows,  | 
463  | 
full_pats_TCs = merge (map pat_of pats) (ListPair.zip (given_pats, TCs)),  | 
|
464  | 
TCs = TCs}  | 
|
465  | 
end;  | 
|
466  | 
||
467  | 
||
468  | 
(*---------------------------------------------------------------------------  | 
|
469  | 
* Perform the extraction without making the definition. Definition and  | 
|
470  | 
* extraction commute for the non-nested case. (Deferred recdefs)  | 
|
471  | 
*  | 
|
472  | 
* The purpose of wfrec_eqns is merely to instantiate the recursion theorem  | 
|
473  | 
* and extract termination conditions: no definition is made.  | 
|
474  | 
*---------------------------------------------------------------------------*)  | 
|
475  | 
||
476  | 
fun wfrec_eqns thy fid tflCongs eqns =  | 
|
477  | 
 let val {lhs,rhs} = S.dest_eq (hd eqns)
 | 
|
478  | 
val (f,args) = S.strip_comb lhs  | 
|
479  | 
val (fname,fty) = dest_atom f  | 
|
480  | 
val (SV,a) = front_last args (* SV = schematic variables *)  | 
|
481  | 
val g = list_comb(f,SV)  | 
|
482  | 
val h = Free(fname,type_of g)  | 
|
483  | 
val eqns1 = map (subst_free[(g,h)]) eqns  | 
|
| 16505 | 484  | 
     val {functional as Abs(x, Ty, _),  pats} = mk_functional thy eqns1
 | 
| 10769 | 485  | 
val given_pats = givens pats  | 
| 16505 | 486  | 
(* val f = Free(x,Ty) *)  | 
| 10769 | 487  | 
     val Type("fun", [f_dty, f_rty]) = Ty
 | 
| 16505 | 488  | 
val dummy = if x<>fid then  | 
| 10769 | 489  | 
raise TFL_ERR "wfrec_eqns"  | 
490  | 
                                      ("Expected a definition of " ^
 | 
|
491  | 
quote fid ^ " but found one of " ^  | 
|
| 16505 | 492  | 
quote x)  | 
| 10769 | 493  | 
else ()  | 
494  | 
val (case_rewrites,context_congs) = extraction_thms thy  | 
|
495  | 
val tych = Thry.typecheck thy  | 
|
496  | 
val WFREC_THM0 = R.ISPEC (tych functional) Thms.WFREC_COROLLARY  | 
|
497  | 
     val Const("All",_) $ Abs(Rname,Rtype,_) = concl WFREC_THM0
 | 
|
| 
20088
 
bffda4cd0f79
replaced Term.variant(list) by Name.variant(_list);
 
wenzelm 
parents: 
20062 
diff
changeset
 | 
498  | 
val R = Free (Name.variant (foldr add_term_names [] eqns) Rname,  | 
| 10769 | 499  | 
Rtype)  | 
500  | 
val WFREC_THM = R.ISPECL [tych R, tych g] WFREC_THM0  | 
|
501  | 
val ([proto_def, WFR],_) = S.strip_imp(concl WFREC_THM)  | 
|
502  | 
val dummy =  | 
|
503  | 
if !trace then  | 
|
504  | 
               writeln ("ORIGINAL PROTO_DEF: " ^
 | 
|
| 22578 | 505  | 
Sign.string_of_term thy proto_def)  | 
| 10769 | 506  | 
else ()  | 
507  | 
val R1 = S.rand WFR  | 
|
508  | 
val corollary' = R.UNDISCH(R.UNDISCH WFREC_THM)  | 
|
509  | 
val corollaries = map (fn pat => R.SPEC (tych pat) corollary') given_pats  | 
|
510  | 
val corollaries' = map (rewrite_rule case_rewrites) corollaries  | 
|
511  | 
fun extract X = R.CONTEXT_REWRITE_RULE  | 
|
512  | 
(f, R1::SV, cut_apply, tflCongs@context_congs) X  | 
|
513  | 
 in {proto_def = proto_def,
 | 
|
514  | 
SV=SV,  | 
|
515  | 
WFR=WFR,  | 
|
516  | 
pats=pats,  | 
|
517  | 
extracta = map extract corollaries'}  | 
|
518  | 
end;  | 
|
519  | 
||
520  | 
||
521  | 
(*---------------------------------------------------------------------------  | 
|
522  | 
* Define the constant after extracting the termination conditions. The  | 
|
523  | 
* wellfounded relation used in the definition is computed by using the  | 
|
524  | 
* choice operator on the extracted conditions (plus the condition that  | 
|
525  | 
* such a relation must be wellfounded).  | 
|
526  | 
*---------------------------------------------------------------------------*)  | 
|
527  | 
||
528  | 
fun lazyR_def thy fid tflCongs eqns =  | 
|
529  | 
 let val {proto_def,WFR,pats,extracta,SV} =
 | 
|
530  | 
wfrec_eqns thy fid tflCongs eqns  | 
|
531  | 
val R1 = S.rand WFR  | 
|
532  | 
val f = #lhs(S.dest_eq proto_def)  | 
|
533  | 
val (extractants,TCl) = ListPair.unzip extracta  | 
|
534  | 
val dummy = if !trace  | 
|
535  | 
then (writeln "Extractants = ";  | 
|
536  | 
prths extractants;  | 
|
537  | 
())  | 
|
538  | 
else ()  | 
|
| 
15574
 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 
skalberg 
parents: 
15570 
diff
changeset
 | 
539  | 
val TCs = foldr (gen_union (op aconv)) [] TCl  | 
| 10769 | 540  | 
val full_rqt = WFR::TCs  | 
541  | 
     val R' = S.mk_select{Bvar=R1, Body=S.list_mk_conj full_rqt}
 | 
|
542  | 
val R'abs = S.rand R'  | 
|
543  | 
val proto_def' = subst_free[(R1,R')] proto_def  | 
|
544  | 
     val dummy = if !trace then writeln ("proto_def' = " ^
 | 
|
545  | 
Sign.string_of_term  | 
|
| 22578 | 546  | 
thy proto_def')  | 
| 10769 | 547  | 
else ()  | 
548  | 
     val {lhs,rhs} = S.dest_eq proto_def'
 | 
|
549  | 
val (c,args) = S.strip_comb lhs  | 
|
| 16505 | 550  | 
val (name,Ty) = dest_atom c  | 
| 22578 | 551  | 
val defn = mk_const_def thy (name, Ty, S.list_mk_abs (args,rhs))  | 
| 18358 | 552  | 
val ([def0], theory) =  | 
| 10769 | 553  | 
thy  | 
554  | 
|> PureThy.add_defs_i false  | 
|
555  | 
[Thm.no_attributes (fid ^ "_def", defn)]  | 
|
| 
19927
 
9286e99b2808
refrain from reforming TFL -- back to previous revision;
 
wenzelm 
parents: 
19925 
diff
changeset
 | 
556  | 
val def = Thm.freezeT def0;  | 
| 10769 | 557  | 
     val dummy = if !trace then writeln ("DEF = " ^ string_of_thm def)
 | 
558  | 
else ()  | 
|
559  | 
(* val fconst = #lhs(S.dest_eq(concl def)) *)  | 
|
560  | 
val tych = Thry.typecheck theory  | 
|
561  | 
val full_rqt_prop = map (Dcterm.mk_prop o tych) full_rqt  | 
|
562  | 
(*lcp: a lot of object-logic inference to remove*)  | 
|
563  | 
val baz = R.DISCH_ALL  | 
|
| 16853 | 564  | 
(fold_rev R.DISCH full_rqt_prop  | 
| 10769 | 565  | 
(R.LIST_CONJ extractants))  | 
566  | 
     val dum = if !trace then writeln ("baz = " ^ string_of_thm baz)
 | 
|
567  | 
else ()  | 
|
568  | 
val f_free = Free (fid, fastype_of f) (*'cos f is a Const*)  | 
|
569  | 
val SV' = map tych SV;  | 
|
570  | 
val SVrefls = map reflexive SV'  | 
|
| 16853 | 571  | 
val def0 = (fold (fn x => fn th => R.rbeta(combination th x))  | 
| 10769 | 572  | 
SVrefls def)  | 
573  | 
RS meta_eq_to_obj_eq  | 
|
574  | 
val def' = R.MP (R.SPEC (tych R') (R.GEN (tych R1) baz)) def0  | 
|
575  | 
val body_th = R.LIST_CONJ (map R.ASSUME full_rqt_prop)  | 
|
| 
11455
 
e07927b980ec
defer_recdef (lazyR_def) now looks for theorem Hilbert_Choice.tfl_some
 
paulson 
parents: 
10769 
diff
changeset
 | 
576  | 
val SELECT_AX = (*in this way we hope to avoid a STATIC dependence upon  | 
| 
 
e07927b980ec
defer_recdef (lazyR_def) now looks for theorem Hilbert_Choice.tfl_some
 
paulson 
parents: 
10769 
diff
changeset
 | 
577  | 
theory Hilbert_Choice*)  | 
| 16505 | 578  | 
thm "Hilbert_Choice.tfl_some"  | 
| 18678 | 579  | 
handle ERROR msg => cat_error msg  | 
| 
11455
 
e07927b980ec
defer_recdef (lazyR_def) now looks for theorem Hilbert_Choice.tfl_some
 
paulson 
parents: 
10769 
diff
changeset
 | 
580  | 
"defer_recdef requires theory Main or at least Hilbert_Choice as parent"  | 
| 
 
e07927b980ec
defer_recdef (lazyR_def) now looks for theorem Hilbert_Choice.tfl_some
 
paulson 
parents: 
10769 
diff
changeset
 | 
581  | 
val bar = R.MP (R.ISPECL[tych R'abs, tych R1] SELECT_AX) body_th  | 
| 10769 | 582  | 
 in {theory = theory, R=R1, SV=SV,
 | 
| 16853 | 583  | 
rules = fold (U.C R.MP) (R.CONJUNCTS bar) def',  | 
| 10769 | 584  | 
full_pats_TCs = merge (map pat_of pats) (ListPair.zip (givens pats, TCl)),  | 
585  | 
patterns = pats}  | 
|
586  | 
end;  | 
|
587  | 
||
588  | 
||
589  | 
||
590  | 
(*----------------------------------------------------------------------------  | 
|
591  | 
*  | 
|
592  | 
* INDUCTION THEOREM  | 
|
593  | 
*  | 
|
594  | 
*---------------------------------------------------------------------------*)  | 
|
595  | 
||
596  | 
||
597  | 
(*------------------------ Miscellaneous function --------------------------  | 
|
598  | 
*  | 
|
599  | 
* [x_1,...,x_n] ?v_1...v_n. M[v_1,...,v_n]  | 
|
600  | 
* -----------------------------------------------------------  | 
|
601  | 
* ( M[x_1,...,x_n], [(x_i,?v_1...v_n. M[v_1,...,v_n]),  | 
|
602  | 
* ...  | 
|
603  | 
* (x_j,?v_n. M[x_1,...,x_(n-1),v_n])] )  | 
|
604  | 
*  | 
|
605  | 
* This function is totally ad hoc. Used in the production of the induction  | 
|
606  | 
* theorem. The nchotomy theorem can have clauses that look like  | 
|
607  | 
*  | 
|
608  | 
* ?v1..vn. z = C vn..v1  | 
|
609  | 
*  | 
|
610  | 
* in which the order of quantification is not the order of occurrence of the  | 
|
611  | 
* quantified variables as arguments to C. Since we have no control over this  | 
|
612  | 
* aspect of the nchotomy theorem, we make the correspondence explicit by  | 
|
613  | 
* pairing the incoming new variable with the term it gets beta-reduced into.  | 
|
614  | 
*---------------------------------------------------------------------------*)  | 
|
615  | 
||
616  | 
fun alpha_ex_unroll (xlist, tm) =  | 
|
617  | 
let val (qvars,body) = S.strip_exists tm  | 
|
618  | 
val vlist = #2(S.strip_comb (S.rhs body))  | 
|
619  | 
val plist = ListPair.zip (vlist, xlist)  | 
|
| 17314 | 620  | 
val args = map (the o AList.lookup (op aconv) plist) qvars  | 
| 15531 | 621  | 
handle Option => sys_error  | 
| 10769 | 622  | 
"TFL fault [alpha_ex_unroll]: no correspondence"  | 
623  | 
fun build ex [] = []  | 
|
624  | 
| build (_$rex) (v::rst) =  | 
|
| 18176 | 625  | 
let val ex1 = Term.betapply(rex, v)  | 
| 10769 | 626  | 
in ex1 :: build ex1 rst  | 
627  | 
end  | 
|
628  | 
val (nex::exl) = rev (tm::build tm args)  | 
|
629  | 
in  | 
|
630  | 
(nex, ListPair.zip (args, rev exl))  | 
|
631  | 
end;  | 
|
632  | 
||
633  | 
||
634  | 
||
635  | 
(*----------------------------------------------------------------------------  | 
|
636  | 
*  | 
|
637  | 
* PROVING COMPLETENESS OF PATTERNS  | 
|
638  | 
*  | 
|
639  | 
*---------------------------------------------------------------------------*)  | 
|
640  | 
||
641  | 
fun mk_case ty_info usednames thy =  | 
|
642  | 
let  | 
|
643  | 
val divide = ipartition (gvvariant usednames)  | 
|
644  | 
val tych = Thry.typecheck thy  | 
|
645  | 
fun tych_binding(x,y) = (tych x, tych y)  | 
|
646  | 
fun fail s = raise TFL_ERR "mk_case" s  | 
|
647  | 
 fun mk{rows=[],...} = fail"no rows"
 | 
|
648  | 
   | mk{path=[], rows = [([], (thm, bindings))]} =
 | 
|
649  | 
R.IT_EXISTS (map tych_binding bindings) thm  | 
|
650  | 
   | mk{path = u::rstp, rows as (p::_, _)::_} =
 | 
|
651  | 
let val (pat_rectangle,rights) = ListPair.unzip rows  | 
|
652  | 
val col0 = map hd pat_rectangle  | 
|
653  | 
val pat_rectangle' = map tl pat_rectangle  | 
|
654  | 
in  | 
|
655  | 
if (forall is_Free col0) (* column 0 is all variables *)  | 
|
656  | 
then let val rights' = map (fn ((thm,theta),v) => (thm,theta@[(u,v)]))  | 
|
657  | 
(ListPair.zip (rights, col0))  | 
|
658  | 
          in mk{path = rstp, rows = ListPair.zip (pat_rectangle', rights')}
 | 
|
659  | 
end  | 
|
660  | 
else (* column 0 is all constructors *)  | 
|
661  | 
let val Type (ty_name,_) = type_of p  | 
|
662  | 
in  | 
|
663  | 
case (ty_info ty_name)  | 
|
| 15531 | 664  | 
     of NONE => fail("Not a known datatype: "^ty_name)
 | 
665  | 
      | SOME{constructors,nchotomy} =>
 | 
|
| 10769 | 666  | 
let val thm' = R.ISPEC (tych u) nchotomy  | 
667  | 
val disjuncts = S.strip_disj (concl thm')  | 
|
668  | 
val subproblems = divide(constructors, rows)  | 
|
669  | 
val groups = map #group subproblems  | 
|
670  | 
and new_formals = map #new_formals subproblems  | 
|
671  | 
val existentials = ListPair.map alpha_ex_unroll  | 
|
672  | 
(new_formals, disjuncts)  | 
|
673  | 
val constraints = map #1 existentials  | 
|
674  | 
val vexl = map #2 existentials  | 
|
675  | 
fun expnd tm (pats,(th,b)) = (pats,(R.SUBS[R.ASSUME(tych tm)]th,b))  | 
|
676  | 
            val news = map (fn (nf,rows,c) => {path = nf@rstp,
 | 
|
677  | 
rows = map (expnd c) rows})  | 
|
678  | 
(U.zip3 new_formals groups constraints)  | 
|
679  | 
val recursive_thms = map mk news  | 
|
| 15570 | 680  | 
val build_exists = Library.foldr  | 
| 10769 | 681  | 
(fn((x,t), th) =>  | 
682  | 
R.CHOOSE (tych x, R.ASSUME (tych t)) th)  | 
|
683  | 
val thms' = ListPair.map build_exists (vexl, recursive_thms)  | 
|
684  | 
val same_concls = R.EVEN_ORS thms'  | 
|
685  | 
in R.DISJ_CASESL thm' same_concls  | 
|
686  | 
end  | 
|
687  | 
end end  | 
|
688  | 
in mk  | 
|
689  | 
end;  | 
|
690  | 
||
691  | 
||
692  | 
fun complete_cases thy =  | 
|
693  | 
let val tych = Thry.typecheck thy  | 
|
694  | 
val ty_info = Thry.induct_info thy  | 
|
695  | 
in fn pats =>  | 
|
| 
15574
 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 
skalberg 
parents: 
15570 
diff
changeset
 | 
696  | 
let val names = foldr add_term_names [] pats  | 
| 10769 | 697  | 
val T = type_of (hd pats)  | 
| 
20088
 
bffda4cd0f79
replaced Term.variant(list) by Name.variant(_list);
 
wenzelm 
parents: 
20062 
diff
changeset
 | 
698  | 
val aname = Name.variant names "a"  | 
| 
 
bffda4cd0f79
replaced Term.variant(list) by Name.variant(_list);
 
wenzelm 
parents: 
20062 
diff
changeset
 | 
699  | 
val vname = Name.variant (aname::names) "v"  | 
| 10769 | 700  | 
val a = Free (aname, T)  | 
701  | 
val v = Free (vname, T)  | 
|
702  | 
val a_eq_v = HOLogic.mk_eq(a,v)  | 
|
703  | 
     val ex_th0 = R.EXISTS (tych (S.mk_exists{Bvar=v,Body=a_eq_v}), tych a)
 | 
|
704  | 
(R.REFL (tych a))  | 
|
705  | 
val th0 = R.ASSUME (tych a_eq_v)  | 
|
706  | 
val rows = map (fn x => ([x], (th0,[]))) pats  | 
|
707  | 
in  | 
|
708  | 
R.GEN (tych a)  | 
|
709  | 
(R.RIGHT_ASSOC  | 
|
710  | 
(R.CHOOSE(tych v, ex_th0)  | 
|
711  | 
(mk_case ty_info (vname::aname::names)  | 
|
712  | 
                 thy {path=[v], rows=rows})))
 | 
|
713  | 
end end;  | 
|
714  | 
||
715  | 
||
716  | 
(*---------------------------------------------------------------------------  | 
|
717  | 
* Constructing induction hypotheses: one for each recursive call.  | 
|
718  | 
*  | 
|
719  | 
* Note. R will never occur as a variable in the ind_clause, because  | 
|
720  | 
* to do so, it would have to be from a nested definition, and we don't  | 
|
721  | 
* allow nested defns to have R variable.  | 
|
722  | 
*  | 
|
723  | 
* Note. When the context is empty, there can be no local variables.  | 
|
724  | 
*---------------------------------------------------------------------------*)  | 
|
725  | 
(*  | 
|
726  | 
local infix 5 ==>  | 
|
727  | 
      fun (tm1 ==> tm2) = S.mk_imp{ant = tm1, conseq = tm2}
 | 
|
728  | 
in  | 
|
729  | 
fun build_ih f P (pat,TCs) =  | 
|
730  | 
let val globals = S.free_vars_lr pat  | 
|
| 15570 | 731  | 
fun nested tm = isSome (S.find_term (curry (op aconv) f) tm)  | 
| 10769 | 732  | 
fun dest_TC tm =  | 
733  | 
let val (cntxt,R_y_pat) = S.strip_imp(#2(S.strip_forall tm))  | 
|
734  | 
val (R,y,_) = S.dest_relation R_y_pat  | 
|
735  | 
val P_y = if (nested tm) then R_y_pat ==> P$y else P$y  | 
|
736  | 
in case cntxt  | 
|
737  | 
of [] => (P_y, (tm,[]))  | 
|
738  | 
| _ => let  | 
|
739  | 
val imp = S.list_mk_conj cntxt ==> P_y  | 
|
740  | 
val lvs = gen_rems (op aconv) (S.free_vars_lr imp, globals)  | 
|
741  | 
val locals = #2(U.pluck (curry (op aconv) P) lvs) handle U.ERR _ => lvs  | 
|
742  | 
in (S.list_mk_forall(locals,imp), (tm,locals)) end  | 
|
743  | 
end  | 
|
744  | 
in case TCs  | 
|
745  | 
of [] => (S.list_mk_forall(globals, P$pat), [])  | 
|
746  | 
| _ => let val (ihs, TCs_locals) = ListPair.unzip(map dest_TC TCs)  | 
|
747  | 
val ind_clause = S.list_mk_conj ihs ==> P$pat  | 
|
748  | 
in (S.list_mk_forall(globals,ind_clause), TCs_locals)  | 
|
749  | 
end  | 
|
750  | 
end  | 
|
751  | 
end;  | 
|
752  | 
*)  | 
|
753  | 
||
754  | 
local infix 5 ==>  | 
|
755  | 
      fun (tm1 ==> tm2) = S.mk_imp{ant = tm1, conseq = tm2}
 | 
|
756  | 
in  | 
|
757  | 
fun build_ih f (P,SV) (pat,TCs) =  | 
|
758  | 
let val pat_vars = S.free_vars_lr pat  | 
|
759  | 
val globals = pat_vars@SV  | 
|
| 15570 | 760  | 
fun nested tm = isSome (S.find_term (curry (op aconv) f) tm)  | 
| 10769 | 761  | 
fun dest_TC tm =  | 
762  | 
let val (cntxt,R_y_pat) = S.strip_imp(#2(S.strip_forall tm))  | 
|
763  | 
val (R,y,_) = S.dest_relation R_y_pat  | 
|
764  | 
val P_y = if (nested tm) then R_y_pat ==> P$y else P$y  | 
|
765  | 
in case cntxt  | 
|
766  | 
of [] => (P_y, (tm,[]))  | 
|
767  | 
| _ => let  | 
|
768  | 
val imp = S.list_mk_conj cntxt ==> P_y  | 
|
| 
20951
 
868120282837
gen_rem(s) abandoned in favour of remove / subtract
 
haftmann 
parents: 
20548 
diff
changeset
 | 
769  | 
val lvs = subtract (op aconv) globals (S.free_vars_lr imp)  | 
| 10769 | 770  | 
val locals = #2(U.pluck (curry (op aconv) P) lvs) handle U.ERR _ => lvs  | 
771  | 
in (S.list_mk_forall(locals,imp), (tm,locals)) end  | 
|
772  | 
end  | 
|
773  | 
in case TCs  | 
|
774  | 
of [] => (S.list_mk_forall(pat_vars, P$pat), [])  | 
|
775  | 
| _ => let val (ihs, TCs_locals) = ListPair.unzip(map dest_TC TCs)  | 
|
776  | 
val ind_clause = S.list_mk_conj ihs ==> P$pat  | 
|
777  | 
in (S.list_mk_forall(pat_vars,ind_clause), TCs_locals)  | 
|
778  | 
end  | 
|
779  | 
end  | 
|
780  | 
end;  | 
|
781  | 
||
782  | 
(*---------------------------------------------------------------------------  | 
|
783  | 
* This function makes good on the promise made in "build_ih".  | 
|
784  | 
*  | 
|
785  | 
* Input is tm = "(!y. R y pat ==> P y) ==> P pat",  | 
|
786  | 
* TCs = TC_1[pat] ... TC_n[pat]  | 
|
787  | 
* thm = ih1 /\ ... /\ ih_n |- ih[pat]  | 
|
788  | 
*---------------------------------------------------------------------------*)  | 
|
789  | 
fun prove_case f thy (tm,TCs_locals,thm) =  | 
|
790  | 
let val tych = Thry.typecheck thy  | 
|
791  | 
val antc = tych(#ant(S.dest_imp tm))  | 
|
792  | 
val thm' = R.SPEC_ALL thm  | 
|
| 15570 | 793  | 
fun nested tm = isSome (S.find_term (curry (op aconv) f) tm)  | 
| 10769 | 794  | 
fun get_cntxt TC = tych(#ant(S.dest_imp(#2(S.strip_forall(concl TC)))))  | 
795  | 
fun mk_ih ((TC,locals),th2,nested) =  | 
|
796  | 
R.GENL (map tych locals)  | 
|
797  | 
(if nested then R.DISCH (get_cntxt TC) th2 handle U.ERR _ => th2  | 
|
798  | 
else if S.is_imp (concl TC) then R.IMP_TRANS TC th2  | 
|
799  | 
else R.MP th2 TC)  | 
|
800  | 
in  | 
|
801  | 
R.DISCH antc  | 
|
802  | 
(if S.is_imp(concl thm') (* recursive calls in this clause *)  | 
|
803  | 
then let val th1 = R.ASSUME antc  | 
|
804  | 
val TCs = map #1 TCs_locals  | 
|
805  | 
val ylist = map (#2 o S.dest_relation o #2 o S.strip_imp o  | 
|
806  | 
#2 o S.strip_forall) TCs  | 
|
807  | 
val TClist = map (fn(TC,lvs) => (R.SPEC_ALL(R.ASSUME(tych TC)),lvs))  | 
|
808  | 
TCs_locals  | 
|
809  | 
val th2list = map (U.C R.SPEC th1 o tych) ylist  | 
|
810  | 
val nlist = map nested TCs  | 
|
811  | 
val triples = U.zip3 TClist th2list nlist  | 
|
812  | 
val Pylist = map mk_ih triples  | 
|
813  | 
in R.MP thm' (R.LIST_CONJ Pylist) end  | 
|
814  | 
else thm')  | 
|
815  | 
end;  | 
|
816  | 
||
817  | 
||
818  | 
(*---------------------------------------------------------------------------  | 
|
819  | 
*  | 
|
820  | 
* x = (v1,...,vn) |- M[x]  | 
|
821  | 
* ---------------------------------------------  | 
|
822  | 
* ?v1 ... vn. x = (v1,...,vn) |- M[x]  | 
|
823  | 
*  | 
|
824  | 
*---------------------------------------------------------------------------*)  | 
|
825  | 
fun LEFT_ABS_VSTRUCT tych thm =  | 
|
826  | 
let fun CHOOSER v (tm,thm) =  | 
|
827  | 
        let val ex_tm = S.mk_exists{Bvar=v,Body=tm}
 | 
|
828  | 
in (ex_tm, R.CHOOSE(tych v, R.ASSUME (tych ex_tm)) thm)  | 
|
829  | 
end  | 
|
| 15570 | 830  | 
val [veq] = List.filter (can S.dest_eq) (#1 (R.dest_thm thm))  | 
| 10769 | 831  | 
      val {lhs,rhs} = S.dest_eq veq
 | 
832  | 
val L = S.free_vars_lr rhs  | 
|
| 16853 | 833  | 
in #2 (fold_rev CHOOSER L (veq,thm)) end;  | 
| 10769 | 834  | 
|
835  | 
||
836  | 
(*----------------------------------------------------------------------------  | 
|
837  | 
* Input : f, R, and [(pat1,TCs1),..., (patn,TCsn)]  | 
|
838  | 
*  | 
|
839  | 
* Instantiates WF_INDUCTION_THM, getting Sinduct and then tries to prove  | 
|
840  | 
* recursion induction (Rinduct) by proving the antecedent of Sinduct from  | 
|
841  | 
* the antecedent of Rinduct.  | 
|
842  | 
*---------------------------------------------------------------------------*)  | 
|
843  | 
fun mk_induction thy {fconst, R, SV, pat_TCs_list} =
 | 
|
844  | 
let val tych = Thry.typecheck thy  | 
|
845  | 
val Sinduction = R.UNDISCH (R.ISPEC (tych R) Thms.WF_INDUCTION_THM)  | 
|
846  | 
val (pats,TCsl) = ListPair.unzip pat_TCs_list  | 
|
847  | 
val case_thm = complete_cases thy pats  | 
|
848  | 
val domain = (type_of o hd) pats  | 
|
| 
20088
 
bffda4cd0f79
replaced Term.variant(list) by Name.variant(_list);
 
wenzelm 
parents: 
20062 
diff
changeset
 | 
849  | 
val Pname = Name.variant (foldr (Library.foldr add_term_names)  | 
| 
15574
 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 
skalberg 
parents: 
15570 
diff
changeset
 | 
850  | 
[] (pats::TCsl)) "P"  | 
| 10769 | 851  | 
val P = Free(Pname, domain --> HOLogic.boolT)  | 
852  | 
val Sinduct = R.SPEC (tych P) Sinduction  | 
|
853  | 
val Sinduct_assumf = S.rand ((#ant o S.dest_imp o concl) Sinduct)  | 
|
854  | 
val Rassums_TCl' = map (build_ih fconst (P,SV)) pat_TCs_list  | 
|
855  | 
val (Rassums,TCl') = ListPair.unzip Rassums_TCl'  | 
|
856  | 
val Rinduct_assum = R.ASSUME (tych (S.list_mk_conj Rassums))  | 
|
| 18176 | 857  | 
val cases = map (fn pat => Term.betapply (Sinduct_assumf, pat)) pats  | 
| 10769 | 858  | 
val tasks = U.zip3 cases TCl' (R.CONJUNCTS Rinduct_assum)  | 
859  | 
val proved_cases = map (prove_case fconst thy) tasks  | 
|
| 
20088
 
bffda4cd0f79
replaced Term.variant(list) by Name.variant(_list);
 
wenzelm 
parents: 
20062 
diff
changeset
 | 
860  | 
val v = Free (Name.variant (foldr add_term_names [] (map concl proved_cases))  | 
| 10769 | 861  | 
"v",  | 
862  | 
domain)  | 
|
863  | 
val vtyped = tych v  | 
|
864  | 
val substs = map (R.SYM o R.ASSUME o tych o (curry HOLogic.mk_eq v)) pats  | 
|
865  | 
val proved_cases1 = ListPair.map (fn (th,th') => R.SUBS[th]th')  | 
|
866  | 
(substs, proved_cases)  | 
|
867  | 
val abs_cases = map (LEFT_ABS_VSTRUCT tych) proved_cases1  | 
|
868  | 
val dant = R.GEN vtyped (R.DISJ_CASESL (R.ISPEC vtyped case_thm) abs_cases)  | 
|
869  | 
val dc = R.MP Sinduct dant  | 
|
870  | 
val Parg_ty = type_of(#Bvar(S.dest_forall(concl dc)))  | 
|
871  | 
val vars = map (gvvariant[Pname]) (S.strip_prod_type Parg_ty)  | 
|
| 16853 | 872  | 
val dc' = fold_rev (R.GEN o tych) vars  | 
| 10769 | 873  | 
(R.SPEC (tych(S.mk_vstruct Parg_ty vars)) dc)  | 
874  | 
in  | 
|
875  | 
R.GEN (tych P) (R.DISCH (tych(concl Rinduct_assum)) dc')  | 
|
876  | 
end  | 
|
877  | 
handle U.ERR _ => raise TFL_ERR "mk_induction" "failed derivation";  | 
|
878  | 
||
879  | 
||
880  | 
||
881  | 
||
882  | 
(*---------------------------------------------------------------------------  | 
|
883  | 
*  | 
|
884  | 
* POST PROCESSING  | 
|
885  | 
*  | 
|
886  | 
*---------------------------------------------------------------------------*)  | 
|
887  | 
||
888  | 
||
889  | 
fun simplify_induction thy hth ind =  | 
|
890  | 
let val tych = Thry.typecheck thy  | 
|
891  | 
val (asl,_) = R.dest_thm ind  | 
|
892  | 
val (_,tc_eq_tc') = R.dest_thm hth  | 
|
893  | 
val tc = S.lhs tc_eq_tc'  | 
|
894  | 
fun loop [] = ind  | 
|
895  | 
| loop (asm::rst) =  | 
|
896  | 
if (can (Thry.match_term thy asm) tc)  | 
|
897  | 
then R.UNDISCH  | 
|
898  | 
(R.MATCH_MP  | 
|
899  | 
(R.MATCH_MP Thms.simp_thm (R.DISCH (tych asm) ind))  | 
|
900  | 
hth)  | 
|
901  | 
else loop rst  | 
|
902  | 
in loop asl  | 
|
903  | 
end;  | 
|
904  | 
||
905  | 
||
906  | 
(*---------------------------------------------------------------------------  | 
|
907  | 
* The termination condition is an antecedent to the rule, and an  | 
|
908  | 
* assumption to the theorem.  | 
|
909  | 
*---------------------------------------------------------------------------*)  | 
|
910  | 
fun elim_tc tcthm (rule,induction) =  | 
|
911  | 
(R.MP rule tcthm, R.PROVE_HYP tcthm induction)  | 
|
912  | 
||
913  | 
||
| 14240 | 914  | 
fun trace_thms s L =  | 
915  | 
if !trace then writeln (cat_lines (s :: map string_of_thm L))  | 
|
916  | 
else ();  | 
|
917  | 
||
918  | 
fun trace_cterms s L =  | 
|
919  | 
if !trace then writeln (cat_lines (s :: map string_of_cterm L))  | 
|
920  | 
else ();;  | 
|
921  | 
||
922  | 
||
| 11632 | 923  | 
fun postprocess strict {wf_tac, terminator, simplifier} theory {rules,induction,TCs} =
 | 
| 10769 | 924  | 
let val tych = Thry.typecheck theory  | 
| 11632 | 925  | 
val prove = R.prove strict;  | 
| 10769 | 926  | 
|
927  | 
(*---------------------------------------------------------------------  | 
|
928  | 
* Attempt to eliminate WF condition. It's the only assumption of rules  | 
|
929  | 
*---------------------------------------------------------------------*)  | 
|
930  | 
val (rules1,induction1) =  | 
|
| 11632 | 931  | 
let val thm = prove(tych(HOLogic.mk_Trueprop  | 
| 10769 | 932  | 
(hd(#1(R.dest_thm rules)))),  | 
933  | 
wf_tac)  | 
|
934  | 
in (R.PROVE_HYP thm rules, R.PROVE_HYP thm induction)  | 
|
935  | 
end handle U.ERR _ => (rules,induction);  | 
|
936  | 
||
937  | 
(*----------------------------------------------------------------------  | 
|
938  | 
* The termination condition (tc) is simplified to |- tc = tc' (there  | 
|
939  | 
* might not be a change!) and then 3 attempts are made:  | 
|
940  | 
*  | 
|
941  | 
* 1. if |- tc = T, then eliminate it with eqT; otherwise,  | 
|
942  | 
* 2. apply the terminator to tc'. If |- tc' = T then eliminate; else  | 
|
943  | 
* 3. replace tc by tc' in both the rules and the induction theorem.  | 
|
944  | 
*---------------------------------------------------------------------*)  | 
|
945  | 
||
946  | 
fun simplify_tc tc (r,ind) =  | 
|
947  | 
let val tc1 = tych tc  | 
|
| 14240 | 948  | 
val _ = trace_cterms "TC before simplification: " [tc1]  | 
| 10769 | 949  | 
val tc_eq = simplifier tc1  | 
| 14240 | 950  | 
val _ = trace_thms "result: " [tc_eq]  | 
| 10769 | 951  | 
in  | 
952  | 
elim_tc (R.MATCH_MP Thms.eqT tc_eq) (r,ind)  | 
|
953  | 
handle U.ERR _ =>  | 
|
954  | 
(elim_tc (R.MATCH_MP(R.MATCH_MP Thms.rev_eq_mp tc_eq)  | 
|
| 11632 | 955  | 
(prove(tych(HOLogic.mk_Trueprop(S.rhs(concl tc_eq))),  | 
| 10769 | 956  | 
terminator)))  | 
957  | 
(r,ind)  | 
|
958  | 
handle U.ERR _ =>  | 
|
959  | 
(R.UNDISCH(R.MATCH_MP (R.MATCH_MP Thms.simp_thm r) tc_eq),  | 
|
960  | 
simplify_induction theory tc_eq ind))  | 
|
961  | 
end  | 
|
962  | 
||
963  | 
(*----------------------------------------------------------------------  | 
|
964  | 
* Nested termination conditions are harder to get at, since they are  | 
|
965  | 
* left embedded in the body of the function (and in induction  | 
|
966  | 
* theorem hypotheses). Our "solution" is to simplify them, and try to  | 
|
967  | 
* prove termination, but leave the application of the resulting theorem  | 
|
968  | 
* to a higher level. So things go much as in "simplify_tc": the  | 
|
969  | 
* termination condition (tc) is simplified to |- tc = tc' (there might  | 
|
970  | 
* not be a change) and then 2 attempts are made:  | 
|
971  | 
*  | 
|
972  | 
* 1. if |- tc = T, then return |- tc; otherwise,  | 
|
973  | 
* 2. apply the terminator to tc'. If |- tc' = T then return |- tc; else  | 
|
974  | 
* 3. return |- tc = tc'  | 
|
975  | 
*---------------------------------------------------------------------*)  | 
|
976  | 
fun simplify_nested_tc tc =  | 
|
977  | 
let val tc_eq = simplifier (tych (#2 (S.strip_forall tc)))  | 
|
978  | 
in  | 
|
979  | 
R.GEN_ALL  | 
|
980  | 
(R.MATCH_MP Thms.eqT tc_eq  | 
|
981  | 
handle U.ERR _ =>  | 
|
982  | 
(R.MATCH_MP(R.MATCH_MP Thms.rev_eq_mp tc_eq)  | 
|
| 11632 | 983  | 
(prove(tych(HOLogic.mk_Trueprop (S.rhs(concl tc_eq))),  | 
| 10769 | 984  | 
terminator))  | 
985  | 
handle U.ERR _ => tc_eq))  | 
|
986  | 
end  | 
|
987  | 
||
988  | 
(*-------------------------------------------------------------------  | 
|
989  | 
* Attempt to simplify the termination conditions in each rule and  | 
|
990  | 
* in the induction theorem.  | 
|
991  | 
*-------------------------------------------------------------------*)  | 
|
992  | 
fun strip_imp tm = if S.is_neg tm then ([],tm) else S.strip_imp tm  | 
|
993  | 
fun loop ([],extras,R,ind) = (rev R, ind, extras)  | 
|
994  | 
| loop ((r,ftcs)::rst, nthms, R, ind) =  | 
|
995  | 
let val tcs = #1(strip_imp (concl r))  | 
|
| 
20951
 
868120282837
gen_rem(s) abandoned in favour of remove / subtract
 
haftmann 
parents: 
20548 
diff
changeset
 | 
996  | 
val extra_tcs = subtract (op aconv) tcs ftcs  | 
| 10769 | 997  | 
val extra_tc_thms = map simplify_nested_tc extra_tcs  | 
| 16853 | 998  | 
val (r1,ind1) = fold simplify_tc tcs (r,ind)  | 
| 10769 | 999  | 
val r2 = R.FILTER_DISCH_ALL(not o S.is_WFR) r1  | 
1000  | 
in loop(rst, nthms@extra_tc_thms, r2::R, ind1)  | 
|
1001  | 
end  | 
|
1002  | 
val rules_tcs = ListPair.zip (R.CONJUNCTS rules1, TCs)  | 
|
1003  | 
val (rules2,ind2,extras) = loop(rules_tcs,[],[],induction1)  | 
|
1004  | 
in  | 
|
1005  | 
  {induction = ind2, rules = R.LIST_CONJ rules2, nested_tcs = extras}
 | 
|
1006  | 
end;  | 
|
1007  | 
||
1008  | 
||
1009  | 
end;  |