author | wenzelm |
Fri, 17 Jul 2009 23:11:40 +0200 | |
changeset 32035 | 8e77b6a250d5 |
parent 31784 | bd3486c57ba3 |
child 32952 | aeb1e44fbc19 |
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 |
|
20344 | 4 |
A package for general recursive function definitions. |
20270
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
krauss
parents:
diff
changeset
|
5 |
|
20344 | 6 |
Automatic splitting of overlapping constructor patterns. This is a preprocessing step which |
20270
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
krauss
parents:
diff
changeset
|
7 |
turns a specification with overlaps into an overlap-free specification. |
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
krauss
parents:
diff
changeset
|
8 |
|
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
krauss
parents:
diff
changeset
|
9 |
*) |
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
krauss
parents:
diff
changeset
|
10 |
|
20344 | 11 |
signature FUNDEF_SPLIT = |
20270
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
krauss
parents:
diff
changeset
|
12 |
sig |
20289 | 13 |
val split_some_equations : |
21237 | 14 |
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
|
15 |
|
20523
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20344
diff
changeset
|
16 |
val split_all_equations : |
21237 | 17 |
Proof.context -> term list -> term list list |
20270
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
krauss
parents:
diff
changeset
|
18 |
end |
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
krauss
parents:
diff
changeset
|
19 |
|
20344 | 20 |
structure FundefSplit : FUNDEF_SPLIT = |
20270
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
krauss
parents:
diff
changeset
|
21 |
struct |
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
krauss
parents:
diff
changeset
|
22 |
|
21051
c49467a9c1e1
Switched function package to use the new package for inductive predicates.
krauss
parents:
20654
diff
changeset
|
23 |
open FundefLib |
20270
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
krauss
parents:
diff
changeset
|
24 |
|
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
krauss
parents:
diff
changeset
|
25 |
(* We use proof context for the variable management *) |
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
krauss
parents:
diff
changeset
|
26 |
(* FIXME: no __ *) |
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
krauss
parents:
diff
changeset
|
27 |
|
20344 | 28 |
fun new_var ctx vs T = |
29 |
let |
|
20270
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
krauss
parents:
diff
changeset
|
30 |
val [v] = Variable.variant_frees ctx vs [("v", T)] |
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
krauss
parents:
diff
changeset
|
31 |
in |
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
krauss
parents:
diff
changeset
|
32 |
(Free v :: vs, Free v) |
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
krauss
parents:
diff
changeset
|
33 |
end |
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
krauss
parents:
diff
changeset
|
34 |
|
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
krauss
parents:
diff
changeset
|
35 |
fun saturate ctx vs t = |
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
krauss
parents:
diff
changeset
|
36 |
fold (fn T => fn (vs, t) => new_var ctx vs T |> apsnd (curry op $ t)) |
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
krauss
parents:
diff
changeset
|
37 |
(binder_types (fastype_of t)) (vs, t) |
21237 | 38 |
|
39 |
||
20270
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
krauss
parents:
diff
changeset
|
40 |
(* This is copied from "fundef_datatype.ML" *) |
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
krauss
parents:
diff
changeset
|
41 |
fun inst_constrs_of thy (T as Type (name, _)) = |
32035 | 42 |
map (fn (Cn,CT) => |
43 |
Envir.subst_term_types (Sign.typ_match thy (body_type CT, T) Vartab.empty) (Const (Cn, CT))) |
|
31784 | 44 |
(the (Datatype.get_constrs thy name)) |
25402 | 45 |
| inst_constrs_of thy T = raise TYPE ("inst_constrs_of", [T], []) |
21237 | 46 |
|
47 |
||
48 |
||
20636 | 49 |
|
50 |
fun join ((vs1,sub1), (vs2,sub2)) = (merge (op aconv) (vs1,vs2), sub1 @ sub2) |
|
25538 | 51 |
fun join_product (xs, ys) = map_product (curry join) xs ys |
20636 | 52 |
|
53 |
fun join_list [] = [] |
|
54 |
| join_list xs = foldr1 (join_product) xs |
|
55 |
||
56 |
||
57 |
exception DISJ |
|
58 |
||
59 |
fun pattern_subtract_subst ctx vs t t' = |
|
20344 | 60 |
let |
20636 | 61 |
exception DISJ |
62 |
fun pattern_subtract_subst_aux vs _ (Free v2) = [] |
|
63 |
| pattern_subtract_subst_aux vs (v as (Free (_, T))) t' = |
|
20344 | 64 |
let |
20636 | 65 |
fun foo constr = |
66 |
let |
|
67 |
val (vs', t) = saturate ctx vs constr |
|
68 |
val substs = pattern_subtract_subst ctx vs' t t' |
|
69 |
in |
|
70 |
map (fn (vs, subst) => (vs, (v,t)::subst)) substs |
|
71 |
end |
|
20270
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
krauss
parents:
diff
changeset
|
72 |
in |
20636 | 73 |
flat (map foo (inst_constrs_of (ProofContext.theory_of ctx) T)) |
74 |
end |
|
75 |
| pattern_subtract_subst_aux vs t t' = |
|
76 |
let |
|
77 |
val (C, ps) = strip_comb t |
|
78 |
val (C', qs) = strip_comb t' |
|
79 |
in |
|
80 |
if C = C' |
|
81 |
then flat (map2 (pattern_subtract_subst_aux vs) ps qs) |
|
82 |
else raise DISJ |
|
20270
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
krauss
parents:
diff
changeset
|
83 |
end |
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
krauss
parents:
diff
changeset
|
84 |
in |
20636 | 85 |
pattern_subtract_subst_aux vs t t' |
86 |
handle DISJ => [(vs, [])] |
|
20270
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
krauss
parents:
diff
changeset
|
87 |
end |
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
krauss
parents:
diff
changeset
|
88 |
|
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
krauss
parents:
diff
changeset
|
89 |
|
20523
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20344
diff
changeset
|
90 |
(* p - q *) |
20270
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
krauss
parents:
diff
changeset
|
91 |
fun pattern_subtract ctx eq2 eq1 = |
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
krauss
parents:
diff
changeset
|
92 |
let |
20636 | 93 |
val thy = ProofContext.theory_of ctx |
21237 | 94 |
|
20523
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20344
diff
changeset
|
95 |
val (vs, feq1 as (_ $ (_ $ lhs1 $ _))) = dest_all_all eq1 |
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20344
diff
changeset
|
96 |
val (_, _ $ (_ $ lhs2 $ _)) = dest_all_all eq2 |
21237 | 97 |
|
20523
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20344
diff
changeset
|
98 |
val substs = pattern_subtract_subst ctx vs lhs1 lhs2 |
21237 | 99 |
|
20523
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20344
diff
changeset
|
100 |
fun instantiate (vs', sigma) = |
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20344
diff
changeset
|
101 |
let |
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20344
diff
changeset
|
102 |
val t = Pattern.rewrite_term thy sigma [] feq1 |
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20344
diff
changeset
|
103 |
in |
27330 | 104 |
fold_rev Logic.all (map Free (frees_in_term ctx t) inter vs') t |
20523
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20344
diff
changeset
|
105 |
end |
20270
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
krauss
parents:
diff
changeset
|
106 |
in |
20654
d80502f0d701
1. Function package accepts a parameter (default "some_term"), which specifies the functions
krauss
parents:
20636
diff
changeset
|
107 |
map instantiate substs |
20270
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
krauss
parents:
diff
changeset
|
108 |
end |
21237 | 109 |
|
20270
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
krauss
parents:
diff
changeset
|
110 |
|
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
krauss
parents:
diff
changeset
|
111 |
(* ps - p' *) |
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
krauss
parents:
diff
changeset
|
112 |
fun pattern_subtract_from_many ctx p'= |
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
krauss
parents:
diff
changeset
|
113 |
flat o map (pattern_subtract ctx p') |
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
krauss
parents:
diff
changeset
|
114 |
|
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
krauss
parents:
diff
changeset
|
115 |
(* in reverse order *) |
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
krauss
parents:
diff
changeset
|
116 |
fun pattern_subtract_many ctx ps' = |
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
krauss
parents:
diff
changeset
|
117 |
fold_rev (pattern_subtract_from_many ctx) ps' |
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
krauss
parents:
diff
changeset
|
118 |
|
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
krauss
parents:
diff
changeset
|
119 |
|
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
krauss
parents:
diff
changeset
|
120 |
|
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
krauss
parents:
diff
changeset
|
121 |
fun split_some_equations ctx eqns = |
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
krauss
parents:
diff
changeset
|
122 |
let |
20338
ecdfc96cf4d0
Added Keywords: "otherwise" and "sequential", needed for function package's
krauss
parents:
20289
diff
changeset
|
123 |
fun split_aux prev [] = [] |
20523
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20344
diff
changeset
|
124 |
| split_aux prev ((true, eq) :: es) = pattern_subtract_many ctx prev [eq] |
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20344
diff
changeset
|
125 |
:: split_aux (eq :: prev) es |
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20344
diff
changeset
|
126 |
| split_aux prev ((false, eq) :: es) = [eq] |
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20344
diff
changeset
|
127 |
:: split_aux (eq :: prev) es |
20270
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
krauss
parents:
diff
changeset
|
128 |
in |
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
krauss
parents:
diff
changeset
|
129 |
split_aux [] eqns |
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
krauss
parents:
diff
changeset
|
130 |
end |
21237 | 131 |
|
20523
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20344
diff
changeset
|
132 |
fun split_all_equations ctx = |
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20344
diff
changeset
|
133 |
split_some_equations ctx o map (pair true) |
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20344
diff
changeset
|
134 |
|
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20344
diff
changeset
|
135 |
|
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20344
diff
changeset
|
136 |
|
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20344
diff
changeset
|
137 |
|
20270
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
krauss
parents:
diff
changeset
|
138 |
end |