| author | wenzelm | 
| Fri, 17 Jun 2011 23:20:34 +0200 | |
| changeset 43426 | 24e2e2f0032b | 
| parent 42483 | 39eefaef816a | 
| child 54406 | a2d18fea844a | 
| permissions | -rw-r--r-- | 
| 31775 | 1 | (* Title: HOL/Tools/Function/pattern_split.ML | 
| 20270 
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
 krauss parents: diff
changeset | 2 | Author: Alexander Krauss, TU Muenchen | 
| 
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
 krauss parents: diff
changeset | 3 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 4 | Fairly ad-hoc pattern splitting. | 
| 20270 
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
 krauss parents: diff
changeset | 5 | *) | 
| 
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
 krauss parents: diff
changeset | 6 | |
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: 
33049diff
changeset | 7 | signature FUNCTION_SPLIT = | 
| 20270 
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
 krauss parents: diff
changeset | 8 | sig | 
| 20289 | 9 | val split_some_equations : | 
| 21237 | 10 | Proof.context -> (bool * term) list -> term list list | 
| 20270 
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
 krauss parents: diff
changeset | 11 | |
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20344diff
changeset | 12 | val split_all_equations : | 
| 21237 | 13 | Proof.context -> term list -> term list list | 
| 20270 
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
 krauss parents: diff
changeset | 14 | end | 
| 
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
 krauss parents: diff
changeset | 15 | |
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: 
33049diff
changeset | 16 | structure Function_Split : FUNCTION_SPLIT = | 
| 20270 
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
 krauss parents: diff
changeset | 17 | struct | 
| 
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
 krauss parents: diff
changeset | 18 | |
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: 
33049diff
changeset | 19 | open Function_Lib | 
| 20270 
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
 krauss parents: diff
changeset | 20 | |
| 42362 
5528970ac6f7
observe firm naming convention ctxt: Proof.context;
 wenzelm parents: 
42361diff
changeset | 21 | fun new_var ctxt vs T = | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 22 | let | 
| 42362 
5528970ac6f7
observe firm naming convention ctxt: Proof.context;
 wenzelm parents: 
42361diff
changeset | 23 |     val [v] = Variable.variant_frees ctxt vs [("v", T)]
 | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 24 | in | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 25 | (Free v :: vs, Free v) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 26 | end | 
| 20270 
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
 krauss parents: diff
changeset | 27 | |
| 42362 
5528970ac6f7
observe firm naming convention ctxt: Proof.context;
 wenzelm parents: 
42361diff
changeset | 28 | fun saturate ctxt vs t = | 
| 
5528970ac6f7
observe firm naming convention ctxt: Proof.context;
 wenzelm parents: 
42361diff
changeset | 29 | fold (fn T => fn (vs, t) => new_var ctxt vs T |> apsnd (curry op $ t)) | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 30 | (binder_types (fastype_of t)) (vs, t) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 31 | |
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 32 | |
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 33 | (* This is copied from "pat_completeness.ML" *) | 
| 20270 
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
 krauss parents: diff
changeset | 34 | fun inst_constrs_of thy (T as Type (name, _)) = | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 35 | map (fn (Cn,CT) => | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 36 | Envir.subst_term_types (Sign.typ_match thy (body_type CT, T) Vartab.empty) (Const (Cn, CT))) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 37 | (the (Datatype.get_constrs thy name)) | 
| 25402 | 38 |   | inst_constrs_of thy T = raise TYPE ("inst_constrs_of", [T], [])
 | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 39 | |
| 20636 | 40 | |
| 41 | fun join ((vs1,sub1), (vs2,sub2)) = (merge (op aconv) (vs1,vs2), sub1 @ sub2) | |
| 25538 | 42 | fun join_product (xs, ys) = map_product (curry join) xs ys | 
| 20636 | 43 | |
| 44 | exception DISJ | |
| 45 | ||
| 42362 
5528970ac6f7
observe firm naming convention ctxt: Proof.context;
 wenzelm parents: 
42361diff
changeset | 46 | fun pattern_subtract_subst ctxt vs t t' = | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 47 | let | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 48 | exception DISJ | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 49 | fun pattern_subtract_subst_aux vs _ (Free v2) = [] | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 50 | | pattern_subtract_subst_aux vs (v as (Free (_, T))) t' = | 
| 20344 | 51 | let | 
| 42362 
5528970ac6f7
observe firm naming convention ctxt: Proof.context;
 wenzelm parents: 
42361diff
changeset | 52 | fun foo constr = | 
| 
5528970ac6f7
observe firm naming convention ctxt: Proof.context;
 wenzelm parents: 
42361diff
changeset | 53 | let | 
| 
5528970ac6f7
observe firm naming convention ctxt: Proof.context;
 wenzelm parents: 
42361diff
changeset | 54 | val (vs', t) = saturate ctxt vs constr | 
| 
5528970ac6f7
observe firm naming convention ctxt: Proof.context;
 wenzelm parents: 
42361diff
changeset | 55 | val substs = pattern_subtract_subst ctxt vs' t t' | 
| 
5528970ac6f7
observe firm naming convention ctxt: Proof.context;
 wenzelm parents: 
42361diff
changeset | 56 | in | 
| 
5528970ac6f7
observe firm naming convention ctxt: Proof.context;
 wenzelm parents: 
42361diff
changeset | 57 | map (fn (vs, subst) => (vs, (v,t)::subst)) substs | 
| 
5528970ac6f7
observe firm naming convention ctxt: Proof.context;
 wenzelm parents: 
42361diff
changeset | 58 | end | 
| 20270 
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
 krauss parents: diff
changeset | 59 | in | 
| 42362 
5528970ac6f7
observe firm naming convention ctxt: Proof.context;
 wenzelm parents: 
42361diff
changeset | 60 | maps foo (inst_constrs_of (Proof_Context.theory_of ctxt) T) | 
| 20636 | 61 | end | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 62 | | pattern_subtract_subst_aux vs t t' = | 
| 42362 
5528970ac6f7
observe firm naming convention ctxt: Proof.context;
 wenzelm parents: 
42361diff
changeset | 63 | let | 
| 
5528970ac6f7
observe firm naming convention ctxt: Proof.context;
 wenzelm parents: 
42361diff
changeset | 64 | val (C, ps) = strip_comb t | 
| 
5528970ac6f7
observe firm naming convention ctxt: Proof.context;
 wenzelm parents: 
42361diff
changeset | 65 | val (C', qs) = strip_comb t' | 
| 
5528970ac6f7
observe firm naming convention ctxt: Proof.context;
 wenzelm parents: 
42361diff
changeset | 66 | in | 
| 
5528970ac6f7
observe firm naming convention ctxt: Proof.context;
 wenzelm parents: 
42361diff
changeset | 67 | if C = C' | 
| 
5528970ac6f7
observe firm naming convention ctxt: Proof.context;
 wenzelm parents: 
42361diff
changeset | 68 | then flat (map2 (pattern_subtract_subst_aux vs) ps qs) | 
| 
5528970ac6f7
observe firm naming convention ctxt: Proof.context;
 wenzelm parents: 
42361diff
changeset | 69 | else raise DISJ | 
| 
5528970ac6f7
observe firm naming convention ctxt: Proof.context;
 wenzelm parents: 
42361diff
changeset | 70 | end | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 71 | in | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 72 | pattern_subtract_subst_aux vs t t' | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 73 | handle DISJ => [(vs, [])] | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 74 | end | 
| 20270 
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
 krauss parents: diff
changeset | 75 | |
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20344diff
changeset | 76 | (* p - q *) | 
| 42362 
5528970ac6f7
observe firm naming convention ctxt: Proof.context;
 wenzelm parents: 
42361diff
changeset | 77 | fun pattern_subtract ctxt eq2 eq1 = | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 78 | let | 
| 42362 
5528970ac6f7
observe firm naming convention ctxt: Proof.context;
 wenzelm parents: 
42361diff
changeset | 79 | val thy = Proof_Context.theory_of ctxt | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 80 | |
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 81 | val (vs, feq1 as (_ $ (_ $ lhs1 $ _))) = dest_all_all eq1 | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 82 | val (_, _ $ (_ $ lhs2 $ _)) = dest_all_all eq2 | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 83 | |
| 42362 
5528970ac6f7
observe firm naming convention ctxt: Proof.context;
 wenzelm parents: 
42361diff
changeset | 84 | val substs = pattern_subtract_subst ctxt vs lhs1 lhs2 | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 85 | |
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 86 | fun instantiate (vs', sigma) = | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 87 | let | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 88 | val t = Pattern.rewrite_term thy sigma [] feq1 | 
| 42483 | 89 | val xs = fold_aterms | 
| 90 | (fn x as Free (a, _) => | |
| 91 | if not (Variable.is_fixed ctxt a) andalso member (op =) vs' x | |
| 92 | then insert (op =) x else I | |
| 93 | | _ => I) t []; | |
| 94 | in fold Logic.all xs t end | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 95 | in | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 96 | map instantiate substs | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 97 | end | 
| 20270 
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
 krauss parents: diff
changeset | 98 | |
| 
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
 krauss parents: diff
changeset | 99 | (* ps - p' *) | 
| 42362 
5528970ac6f7
observe firm naming convention ctxt: Proof.context;
 wenzelm parents: 
42361diff
changeset | 100 | fun pattern_subtract_from_many ctxt p'= | 
| 
5528970ac6f7
observe firm naming convention ctxt: Proof.context;
 wenzelm parents: 
42361diff
changeset | 101 | maps (pattern_subtract ctxt p') | 
| 20270 
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
 krauss parents: diff
changeset | 102 | |
| 
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
 krauss parents: diff
changeset | 103 | (* in reverse order *) | 
| 42362 
5528970ac6f7
observe firm naming convention ctxt: Proof.context;
 wenzelm parents: 
42361diff
changeset | 104 | fun pattern_subtract_many ctxt ps' = | 
| 
5528970ac6f7
observe firm naming convention ctxt: Proof.context;
 wenzelm parents: 
42361diff
changeset | 105 | fold_rev (pattern_subtract_from_many ctxt) ps' | 
| 20270 
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
 krauss parents: diff
changeset | 106 | |
| 42362 
5528970ac6f7
observe firm naming convention ctxt: Proof.context;
 wenzelm parents: 
42361diff
changeset | 107 | fun split_some_equations ctxt eqns = | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 108 | let | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 109 | fun split_aux prev [] = [] | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 110 | | split_aux prev ((true, eq) :: es) = | 
| 42362 
5528970ac6f7
observe firm naming convention ctxt: Proof.context;
 wenzelm parents: 
42361diff
changeset | 111 | pattern_subtract_many ctxt prev [eq] :: split_aux (eq :: prev) es | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 112 | | split_aux prev ((false, eq) :: es) = | 
| 42362 
5528970ac6f7
observe firm naming convention ctxt: Proof.context;
 wenzelm parents: 
42361diff
changeset | 113 | [eq] :: split_aux (eq :: prev) es | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 114 | in | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 115 | split_aux [] eqns | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 116 | end | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 117 | |
| 42362 
5528970ac6f7
observe firm naming convention ctxt: Proof.context;
 wenzelm parents: 
42361diff
changeset | 118 | fun split_all_equations ctxt = | 
| 
5528970ac6f7
observe firm naming convention ctxt: Proof.context;
 wenzelm parents: 
42361diff
changeset | 119 | split_some_equations ctxt o map (pair true) | 
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20344diff
changeset | 120 | |
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20344diff
changeset | 121 | |
| 20270 
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
 krauss parents: diff
changeset | 122 | end |