| author | haftmann | 
| Tue, 31 Oct 2006 09:28:54 +0100 | |
| changeset 21111 | 624ed9c7c4fe | 
| parent 21022 | 3634641f9405 | 
| child 22271 | 51a80e238b29 | 
| permissions | -rw-r--r-- | 
| 12453 | 1  | 
(* Title: HOL/inductive_codegen.ML  | 
| 11537 | 2  | 
ID: $Id$  | 
| 11539 | 3  | 
Author: Stefan Berghofer, TU Muenchen  | 
| 11537 | 4  | 
|
| 11539 | 5  | 
Code generator for inductive predicates.  | 
| 11537 | 6  | 
*)  | 
7  | 
||
8  | 
signature INDUCTIVE_CODEGEN =  | 
|
9  | 
sig  | 
|
| 18728 | 10  | 
val add : string option -> attribute  | 
| 18708 | 11  | 
val setup : theory -> theory  | 
| 11537 | 12  | 
end;  | 
13  | 
||
14  | 
structure InductiveCodegen : INDUCTIVE_CODEGEN =  | 
|
15  | 
struct  | 
|
16  | 
||
17  | 
open Codegen;  | 
|
18  | 
||
| 12557 | 19  | 
(**** theory data ****)  | 
20  | 
||
| 
18930
 
29d39c10822e
replaced Symtab.merge_multi by local merge_rules;
 
wenzelm 
parents: 
18928 
diff
changeset
 | 
21  | 
fun merge_rules tabs =  | 
| 
 
29d39c10822e
replaced Symtab.merge_multi by local merge_rules;
 
wenzelm 
parents: 
18928 
diff
changeset
 | 
22  | 
Symtab.join (fn _ => fn (ths, ths') =>  | 
| 19025 | 23  | 
gen_merge_lists (Drule.eq_thm_prop o pairself fst) ths ths') tabs;  | 
| 
18930
 
29d39c10822e
replaced Symtab.merge_multi by local merge_rules;
 
wenzelm 
parents: 
18928 
diff
changeset
 | 
24  | 
|
| 16424 | 25  | 
structure CodegenData = TheoryDataFun  | 
26  | 
(struct  | 
|
| 12557 | 27  | 
val name = "HOL/inductive_codegen";  | 
| 
15031
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
28  | 
type T =  | 
| 16645 | 29  | 
    {intros : (thm * string) list Symtab.table,
 | 
| 
15031
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
30  | 
graph : unit Graph.T,  | 
| 16645 | 31  | 
eqns : (thm * string) list Symtab.table};  | 
| 
15031
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
32  | 
val empty =  | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
33  | 
    {intros = Symtab.empty, graph = Graph.empty, eqns = Symtab.empty};
 | 
| 12557 | 34  | 
val copy = I;  | 
| 16424 | 35  | 
val extend = I;  | 
36  | 
  fun merge _ ({intros=intros1, graph=graph1, eqns=eqns1},
 | 
|
| 
15031
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
37  | 
    {intros=intros2, graph=graph2, eqns=eqns2}) =
 | 
| 
18930
 
29d39c10822e
replaced Symtab.merge_multi by local merge_rules;
 
wenzelm 
parents: 
18928 
diff
changeset
 | 
38  | 
    {intros = merge_rules (intros1, intros2),
 | 
| 
15031
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
39  | 
graph = Graph.merge (K true) (graph1, graph2),  | 
| 
18930
 
29d39c10822e
replaced Symtab.merge_multi by local merge_rules;
 
wenzelm 
parents: 
18928 
diff
changeset
 | 
40  | 
eqns = merge_rules (eqns1, eqns2)};  | 
| 12557 | 41  | 
fun print _ _ = ();  | 
| 16424 | 42  | 
end);  | 
| 12557 | 43  | 
|
44  | 
||
45  | 
fun warn thm = warning ("InductiveCodegen: Not a proper clause:\n" ^
 | 
|
46  | 
string_of_thm thm);  | 
|
| 11537 | 47  | 
|
| 
14162
 
f05f299512e9
Fixed problem with "code ind" attribute that caused code generator to
 
berghofe 
parents: 
13105 
diff
changeset
 | 
48  | 
fun add_node (g, x) = Graph.new_node (x, ()) g handle Graph.DUP _ => g;  | 
| 
 
f05f299512e9
Fixed problem with "code ind" attribute that caused code generator to
 
berghofe 
parents: 
13105 
diff
changeset
 | 
49  | 
|
| 20897 | 50  | 
fun add optmod = Thm.declaration_attribute (fn thm => Context.mapping (fn thy =>  | 
| 16645 | 51  | 
let  | 
52  | 
    val {intros, graph, eqns} = CodegenData.get thy;
 | 
|
53  | 
fun thyname_of s = (case optmod of  | 
|
54  | 
NONE => thyname_of_const s thy | SOME s => s);  | 
|
| 12557 | 55  | 
in (case concl_of thm of  | 
56  | 
      _ $ (Const ("op :", _) $ _ $ t) => (case head_of t of
 | 
|
| 
14162
 
f05f299512e9
Fixed problem with "code ind" attribute that caused code generator to
 
berghofe 
parents: 
13105 
diff
changeset
 | 
57  | 
Const (s, _) =>  | 
| 
15574
 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 
skalberg 
parents: 
15570 
diff
changeset
 | 
58  | 
let val cs = foldr add_term_consts [] (prems_of thm)  | 
| 18728 | 59  | 
in CodegenData.put  | 
| 17261 | 60  | 
            {intros = intros |>
 | 
| 
18930
 
29d39c10822e
replaced Symtab.merge_multi by local merge_rules;
 
wenzelm 
parents: 
18928 
diff
changeset
 | 
61  | 
Symtab.update (s, Symtab.lookup_list intros s @ [(thm, thyname_of s)]),  | 
| 
15574
 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 
skalberg 
parents: 
15570 
diff
changeset
 | 
62  | 
graph = foldr (uncurry (Graph.add_edge o pair s))  | 
| 
 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 
skalberg 
parents: 
15570 
diff
changeset
 | 
63  | 
(Library.foldl add_node (graph, s :: cs)) cs,  | 
| 18728 | 64  | 
eqns = eqns} thy  | 
| 
14162
 
f05f299512e9
Fixed problem with "code ind" attribute that caused code generator to
 
berghofe 
parents: 
13105 
diff
changeset
 | 
65  | 
end  | 
| 18728 | 66  | 
| _ => (warn thm; thy))  | 
| 
15031
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
67  | 
    | _ $ (Const ("op =", _) $ t $ _) => (case head_of t of
 | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
68  | 
Const (s, _) =>  | 
| 18728 | 69  | 
          CodegenData.put {intros = intros, graph = graph,
 | 
| 17412 | 70  | 
eqns = eqns |> Symtab.update  | 
| 
18930
 
29d39c10822e
replaced Symtab.merge_multi by local merge_rules;
 
wenzelm 
parents: 
18928 
diff
changeset
 | 
71  | 
(s, Symtab.lookup_list eqns s @ [(thm, thyname_of s)])} thy  | 
| 18728 | 72  | 
| _ => (warn thm; thy))  | 
73  | 
| _ => (warn thm; thy))  | 
|
| 20897 | 74  | 
end) I);  | 
| 12557 | 75  | 
|
76  | 
fun get_clauses thy s =  | 
|
| 
15031
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
77  | 
  let val {intros, graph, ...} = CodegenData.get thy
 | 
| 17412 | 78  | 
in case Symtab.lookup intros s of  | 
| 
21022
 
3634641f9405
Moved old inductive package to old_inductive_package.ML
 
berghofe 
parents: 
20897 
diff
changeset
 | 
79  | 
NONE => (case OldInductivePackage.get_inductive thy s of  | 
| 15531 | 80  | 
NONE => NONE  | 
| 16645 | 81  | 
      | SOME ({names, ...}, {intrs, ...}) =>
 | 
82  | 
SOME (names, thyname_of_const s thy,  | 
|
83  | 
preprocess thy intrs))  | 
|
| 15531 | 84  | 
| SOME _ =>  | 
| 16645 | 85  | 
let  | 
86  | 
val SOME names = find_first  | 
|
87  | 
(fn xs => s mem xs) (Graph.strong_conn graph);  | 
|
88  | 
val intrs = List.concat (map  | 
|
| 17412 | 89  | 
(fn s => the (Symtab.lookup intros s)) names);  | 
| 16645 | 90  | 
val (_, (_, thyname)) = split_last intrs  | 
91  | 
in SOME (names, thyname, preprocess thy (map fst intrs)) end  | 
|
| 
14162
 
f05f299512e9
Fixed problem with "code ind" attribute that caused code generator to
 
berghofe 
parents: 
13105 
diff
changeset
 | 
92  | 
end;  | 
| 12557 | 93  | 
|
94  | 
||
95  | 
(**** improper tuples ****)  | 
|
| 11537 | 96  | 
|
97  | 
fun prod_factors p (Const ("Pair", _) $ t $ u) =
 | 
|
98  | 
p :: prod_factors (1::p) t @ prod_factors (2::p) u  | 
|
99  | 
| prod_factors p _ = [];  | 
|
100  | 
||
101  | 
fun split_prod p ps t = if p mem ps then (case t of  | 
|
102  | 
       Const ("Pair", _) $ t $ u =>
 | 
|
103  | 
split_prod (1::p) ps t @ split_prod (2::p) ps u  | 
|
104  | 
| _ => error "Inconsistent use of products") else [t];  | 
|
105  | 
||
| 
15031
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
106  | 
fun full_split_prod (Const ("Pair", _) $ t $ u) =
 | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
107  | 
full_split_prod t @ full_split_prod u  | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
108  | 
| full_split_prod t = [t];  | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
109  | 
|
| 12557 | 110  | 
datatype factors = FVar of int list list | FFix of int list list;  | 
111  | 
||
112  | 
exception Factors;  | 
|
113  | 
||
114  | 
fun mg_factor (FVar f) (FVar f') = FVar (f inter f')  | 
|
115  | 
| mg_factor (FVar f) (FFix f') =  | 
|
116  | 
if f' subset f then FFix f' else raise Factors  | 
|
117  | 
| mg_factor (FFix f) (FVar f') =  | 
|
118  | 
if f subset f' then FFix f else raise Factors  | 
|
119  | 
| mg_factor (FFix f) (FFix f') =  | 
|
120  | 
if f subset f' andalso f' subset f then FFix f else raise Factors;  | 
|
121  | 
||
122  | 
fun dest_factors (FVar f) = f  | 
|
123  | 
| dest_factors (FFix f) = f;  | 
|
124  | 
||
125  | 
fun infer_factors sg extra_fs (fs, (optf, t)) =  | 
|
126  | 
let fun err s = error (s ^ "\n" ^ Sign.string_of_term sg t)  | 
|
127  | 
in (case (optf, strip_comb t) of  | 
|
| 15531 | 128  | 
(SOME f, (Const (name, _), args)) =>  | 
| 17521 | 129  | 
(case AList.lookup (op =) extra_fs name of  | 
130  | 
NONE => AList.update (op =) (name, getOpt  | 
|
131  | 
(Option.map (mg_factor f) (AList.lookup (op =) fs name), f)) fs  | 
|
| 15531 | 132  | 
| SOME (fs', f') => (mg_factor f (FFix f');  | 
| 15570 | 133  | 
Library.foldl (infer_factors sg extra_fs)  | 
134  | 
(fs, map (Option.map FFix) fs' ~~ args)))  | 
|
| 15531 | 135  | 
| (SOME f, (Var ((name, _), _), [])) =>  | 
| 17521 | 136  | 
AList.update (op =) (name, getOpt  | 
137  | 
(Option.map (mg_factor f) (AList.lookup (op =) fs name), f)) fs  | 
|
| 15531 | 138  | 
| (NONE, _) => fs  | 
| 12557 | 139  | 
| _ => err "Illegal term")  | 
140  | 
handle Factors => err "Product factor mismatch in"  | 
|
141  | 
end;  | 
|
142  | 
||
| 11537 | 143  | 
fun string_of_factors p ps = if p mem ps then  | 
144  | 
    "(" ^ string_of_factors (1::p) ps ^ ", " ^ string_of_factors (2::p) ps ^ ")"
 | 
|
145  | 
else "_";  | 
|
146  | 
||
| 12557 | 147  | 
|
| 11537 | 148  | 
(**** check if a term contains only constructor functions ****)  | 
149  | 
||
150  | 
fun is_constrt thy =  | 
|
151  | 
let  | 
|
| 15570 | 152  | 
val cnstrs = List.concat (List.concat (map  | 
| 11537 | 153  | 
(map (fn (_, (_, _, cs)) => map (apsnd length) cs) o #descr o snd)  | 
154  | 
(Symtab.dest (DatatypePackage.get_datatypes thy))));  | 
|
155  | 
fun check t = (case strip_comb t of  | 
|
156  | 
(Var _, []) => true  | 
|
| 17521 | 157  | 
| (Const (s, _), ts) => (case AList.lookup (op =) cnstrs s of  | 
| 15531 | 158  | 
NONE => false  | 
159  | 
| SOME i => length ts = i andalso forall check ts)  | 
|
| 11537 | 160  | 
| _ => false)  | 
161  | 
in check end;  | 
|
162  | 
||
163  | 
(**** check if a type is an equality type (i.e. doesn't contain fun) ****)  | 
|
164  | 
||
165  | 
fun is_eqT (Type (s, Ts)) = s <> "fun" andalso forall is_eqT Ts  | 
|
166  | 
| is_eqT _ = true;  | 
|
167  | 
||
168  | 
(**** mode inference ****)  | 
|
169  | 
||
| 
15204
 
d15f85347fcb
- Inserted additional check for equality types in check_mode_clause that
 
berghofe 
parents: 
15126 
diff
changeset
 | 
170  | 
fun string_of_mode (iss, is) = space_implode " -> " (map  | 
| 15531 | 171  | 
(fn NONE => "X"  | 
172  | 
| SOME js => enclose "[" "]" (commas (map string_of_int js)))  | 
|
173  | 
(iss @ [SOME is]));  | 
|
| 
15204
 
d15f85347fcb
- Inserted additional check for equality types in check_mode_clause that
 
berghofe 
parents: 
15126 
diff
changeset
 | 
174  | 
|
| 
 
d15f85347fcb
- Inserted additional check for equality types in check_mode_clause that
 
berghofe 
parents: 
15126 
diff
changeset
 | 
175  | 
fun print_modes modes = message ("Inferred modes:\n" ^
 | 
| 
 
d15f85347fcb
- Inserted additional check for equality types in check_mode_clause that
 
berghofe 
parents: 
15126 
diff
changeset
 | 
176  | 
space_implode "\n" (map (fn (s, ms) => s ^ ": " ^ commas (map  | 
| 
 
d15f85347fcb
- Inserted additional check for equality types in check_mode_clause that
 
berghofe 
parents: 
15126 
diff
changeset
 | 
177  | 
string_of_mode ms)) modes));  | 
| 
 
d15f85347fcb
- Inserted additional check for equality types in check_mode_clause that
 
berghofe 
parents: 
15126 
diff
changeset
 | 
178  | 
|
| 11537 | 179  | 
val term_vs = map (fst o fst o dest_Var) o term_vars;  | 
| 
19046
 
bc5c6c9b114e
removed distinct, renamed gen_distinct to distinct;
 
wenzelm 
parents: 
19025 
diff
changeset
 | 
180  | 
val terms_vs = distinct (op =) o List.concat o (map term_vs);  | 
| 11537 | 181  | 
|
182  | 
(** collect all Vars in a term (with duplicates!) **)  | 
|
| 16861 | 183  | 
fun term_vTs tm =  | 
184  | 
fold_aterms (fn Var ((x, _), T) => cons (x, T) | _ => I) tm [];  | 
|
| 11537 | 185  | 
|
186  | 
fun get_args _ _ [] = ([], [])  | 
|
187  | 
| get_args is i (x::xs) = (if i mem is then apfst else apsnd) (cons x)  | 
|
188  | 
(get_args is (i+1) xs);  | 
|
189  | 
||
190  | 
fun merge xs [] = xs  | 
|
191  | 
| merge [] ys = ys  | 
|
192  | 
| merge (x::xs) (y::ys) = if length x >= length y then x::merge xs (y::ys)  | 
|
193  | 
else y::merge (x::xs) ys;  | 
|
194  | 
||
195  | 
fun subsets i j = if i <= j then  | 
|
196  | 
let val is = subsets (i+1) j  | 
|
197  | 
in merge (map (fn ks => i::ks) is) is end  | 
|
198  | 
else [[]];  | 
|
199  | 
||
| 12557 | 200  | 
fun cprod ([], ys) = []  | 
201  | 
| cprod (x :: xs, ys) = map (pair x) ys @ cprod (xs, ys);  | 
|
202  | 
||
| 
15574
 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 
skalberg 
parents: 
15570 
diff
changeset
 | 
203  | 
fun cprods xss = foldr (map op :: o cprod) [[]] xss;  | 
| 12557 | 204  | 
|
205  | 
datatype mode = Mode of (int list option list * int list) * mode option list;  | 
|
206  | 
||
207  | 
fun modes_of modes t =  | 
|
208  | 
let  | 
|
| 15570 | 209  | 
fun mk_modes name args = List.concat  | 
| 12557 | 210  | 
(map (fn (m as (iss, is)) => map (Mode o pair m) (cprods (map  | 
| 15531 | 211  | 
(fn (NONE, _) => [NONE]  | 
212  | 
| (SOME js, arg) => map SOME  | 
|
| 15570 | 213  | 
(List.filter (fn Mode ((_, js'), _) => js=js') (modes_of modes arg)))  | 
| 17521 | 214  | 
(iss ~~ args)))) ((the o AList.lookup (op =) modes) name))  | 
| 12557 | 215  | 
|
216  | 
in (case strip_comb t of  | 
|
| 
14163
 
5ffa75ce4919
Improved handling of modes for equality predicate, to avoid ill-typed
 
berghofe 
parents: 
14162 
diff
changeset
 | 
217  | 
      (Const ("op =", Type (_, [T, _])), _) =>
 | 
| 
 
5ffa75ce4919
Improved handling of modes for equality predicate, to avoid ill-typed
 
berghofe 
parents: 
14162 
diff
changeset
 | 
218  | 
[Mode (([], [1]), []), Mode (([], [2]), [])] @  | 
| 
 
5ffa75ce4919
Improved handling of modes for equality predicate, to avoid ill-typed
 
berghofe 
parents: 
14162 
diff
changeset
 | 
219  | 
(if is_eqT T then [Mode (([], [1, 2]), [])] else [])  | 
| 
 
5ffa75ce4919
Improved handling of modes for equality predicate, to avoid ill-typed
 
berghofe 
parents: 
14162 
diff
changeset
 | 
220  | 
| (Const (name, _), args) => mk_modes name args  | 
| 13038 | 221  | 
| (Var ((name, _), _), args) => mk_modes name args  | 
222  | 
| (Free (name, _), args) => mk_modes name args)  | 
|
| 12557 | 223  | 
end;  | 
224  | 
||
225  | 
datatype indprem = Prem of term list * term | Sidecond of term;  | 
|
226  | 
||
| 11537 | 227  | 
fun select_mode_prem thy modes vs ps =  | 
| 19861 | 228  | 
find_first (is_some o snd) (ps ~~ map  | 
| 12557 | 229  | 
(fn Prem (us, t) => find_first (fn Mode ((_, is), _) =>  | 
| 11537 | 230  | 
let  | 
| 
15031
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
231  | 
val (in_ts, out_ts) = get_args is 1 us;  | 
| 15570 | 232  | 
val (out_ts', in_ts') = List.partition (is_constrt thy) out_ts;  | 
233  | 
val vTs = List.concat (map term_vTs out_ts');  | 
|
| 18964 | 234  | 
val dupTs = map snd (duplicates (op =) vTs) @  | 
| 17521 | 235  | 
List.mapPartial (AList.lookup (op =) vTs) vs;  | 
| 11537 | 236  | 
in  | 
| 
15031
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
237  | 
terms_vs (in_ts @ in_ts') subset vs andalso  | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
238  | 
forall (is_eqT o fastype_of) in_ts' andalso  | 
| 12557 | 239  | 
term_vs t subset vs andalso  | 
| 11537 | 240  | 
forall is_eqT dupTs  | 
241  | 
end)  | 
|
| 15660 | 242  | 
(modes_of modes t handle Option => [Mode (([], []), [])])  | 
| 15531 | 243  | 
| Sidecond t => if term_vs t subset vs then SOME (Mode (([], []), []))  | 
244  | 
else NONE) ps);  | 
|
| 11537 | 245  | 
|
| 12557 | 246  | 
fun check_mode_clause thy arg_vs modes (iss, is) (ts, ps) =  | 
| 11537 | 247  | 
let  | 
| 15570 | 248  | 
val modes' = modes @ List.mapPartial  | 
| 15531 | 249  | 
(fn (_, NONE) => NONE | (v, SOME js) => SOME (v, [([], js)]))  | 
| 12557 | 250  | 
(arg_vs ~~ iss);  | 
| 15531 | 251  | 
fun check_mode_prems vs [] = SOME vs  | 
| 12557 | 252  | 
| check_mode_prems vs ps = (case select_mode_prem thy modes' vs ps of  | 
| 15531 | 253  | 
NONE => NONE  | 
254  | 
| SOME (x, _) => check_mode_prems  | 
|
| 12557 | 255  | 
(case x of Prem (us, _) => vs union terms_vs us | _ => vs)  | 
| 11537 | 256  | 
(filter_out (equal x) ps));  | 
| 15570 | 257  | 
val (in_ts, in_ts') = List.partition (is_constrt thy) (fst (get_args is 1 ts));  | 
| 11537 | 258  | 
val in_vs = terms_vs in_ts;  | 
259  | 
val concl_vs = terms_vs ts  | 
|
260  | 
in  | 
|
| 18964 | 261  | 
forall is_eqT (map snd (duplicates (op =) (List.concat (map term_vTs in_ts)))) andalso  | 
| 
15204
 
d15f85347fcb
- Inserted additional check for equality types in check_mode_clause that
 
berghofe 
parents: 
15126 
diff
changeset
 | 
262  | 
forall (is_eqT o fastype_of) in_ts' andalso  | 
| 11537 | 263  | 
(case check_mode_prems (arg_vs union in_vs) ps of  | 
| 15531 | 264  | 
NONE => false  | 
265  | 
| SOME vs => concl_vs subset vs)  | 
|
| 11537 | 266  | 
end;  | 
267  | 
||
268  | 
fun check_modes_pred thy arg_vs preds modes (p, ms) =  | 
|
| 17521 | 269  | 
let val SOME rs = AList.lookup (op =) preds p  | 
| 15570 | 270  | 
in (p, List.filter (fn m => case find_index  | 
| 
15204
 
d15f85347fcb
- Inserted additional check for equality types in check_mode_clause that
 
berghofe 
parents: 
15126 
diff
changeset
 | 
271  | 
(not o check_mode_clause thy arg_vs modes m) rs of  | 
| 
 
d15f85347fcb
- Inserted additional check for equality types in check_mode_clause that
 
berghofe 
parents: 
15126 
diff
changeset
 | 
272  | 
~1 => true  | 
| 
 
d15f85347fcb
- Inserted additional check for equality types in check_mode_clause that
 
berghofe 
parents: 
15126 
diff
changeset
 | 
273  | 
    | i => (message ("Clause " ^ string_of_int (i+1) ^ " of " ^
 | 
| 
 
d15f85347fcb
- Inserted additional check for equality types in check_mode_clause that
 
berghofe 
parents: 
15126 
diff
changeset
 | 
274  | 
p ^ " violates mode " ^ string_of_mode m); false)) ms)  | 
| 
 
d15f85347fcb
- Inserted additional check for equality types in check_mode_clause that
 
berghofe 
parents: 
15126 
diff
changeset
 | 
275  | 
end;  | 
| 11537 | 276  | 
|
277  | 
fun fixp f x =  | 
|
278  | 
let val y = f x  | 
|
279  | 
in if x = y then x else fixp f y end;  | 
|
280  | 
||
| 12557 | 281  | 
fun infer_modes thy extra_modes factors arg_vs preds = fixp (fn modes =>  | 
| 11537 | 282  | 
map (check_modes_pred thy arg_vs preds (modes @ extra_modes)) modes)  | 
| 12557 | 283  | 
(map (fn (s, (fs, f)) => (s, cprod (cprods (map  | 
| 15531 | 284  | 
(fn NONE => [NONE]  | 
285  | 
| SOME f' => map SOME (subsets 1 (length f' + 1))) fs),  | 
|
| 12557 | 286  | 
subsets 1 (length f + 1)))) factors);  | 
| 11537 | 287  | 
|
288  | 
(**** code generation ****)  | 
|
289  | 
||
290  | 
fun mk_eq (x::xs) =  | 
|
291  | 
let fun mk_eqs _ [] = []  | 
|
292  | 
| mk_eqs a (b::cs) = Pretty.str (a ^ " = " ^ b) :: mk_eqs b cs  | 
|
293  | 
in mk_eqs x xs end;  | 
|
294  | 
||
295  | 
fun mk_tuple xs = Pretty.block (Pretty.str "(" ::
 | 
|
| 15570 | 296  | 
List.concat (separate [Pretty.str ",", Pretty.brk 1] (map single xs)) @  | 
| 11537 | 297  | 
[Pretty.str ")"]);  | 
298  | 
||
| 
15031
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
299  | 
(* convert nested pairs to n-tuple *)  | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
300  | 
|
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
301  | 
fun conv_ntuple [_] t ps = ps  | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
302  | 
| conv_ntuple [_, _] t ps = ps  | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
303  | 
| conv_ntuple us t ps =  | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
304  | 
let  | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
305  | 
        val xs = map (fn i => Pretty.str ("x" ^ string_of_int i))
 | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
306  | 
(1 upto length us);  | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
307  | 
fun ntuple (ys as (x, T) :: xs) U =  | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
308  | 
if T = U then (x, xs)  | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
309  | 
else  | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
310  | 
let  | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
311  | 
              val Type ("*", [U1, U2]) = U;
 | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
312  | 
val (p1, ys1) = ntuple ys U1;  | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
313  | 
val (p2, ys2) = ntuple ys1 U2  | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
314  | 
in (mk_tuple [p1, p2], ys2) end  | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
315  | 
in  | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
316  | 
[Pretty.str "Seq.map (fn", Pretty.brk 1,  | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
317  | 
fst (ntuple (xs ~~ map fastype_of us) (HOLogic.dest_setT (fastype_of t))),  | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
318  | 
Pretty.str " =>", Pretty.brk 1, mk_tuple xs, Pretty.str ")",  | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
319  | 
Pretty.brk 1, parens (Pretty.block ps)]  | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
320  | 
end;  | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
321  | 
|
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
322  | 
(* convert n-tuple to nested pairs *)  | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
323  | 
|
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
324  | 
fun conv_ntuple' fs T ps =  | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
325  | 
let  | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
326  | 
    fun mk_x i = Pretty.str ("x" ^ string_of_int i);
 | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
327  | 
    fun conv i js (Type ("*", [T, U])) =
 | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
328  | 
if js mem fs then  | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
329  | 
let  | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
330  | 
val (p, i') = conv i (1::js) T;  | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
331  | 
val (q, i'') = conv i' (2::js) U  | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
332  | 
in (mk_tuple [p, q], i'') end  | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
333  | 
else (mk_x i, i+1)  | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
334  | 
| conv i js _ = (mk_x i, i+1)  | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
335  | 
val (p, i) = conv 1 [] T  | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
336  | 
in  | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
337  | 
if i > 3 then  | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
338  | 
[Pretty.str "Seq.map (fn", Pretty.brk 1,  | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
339  | 
mk_tuple (map mk_x (1 upto i-1)), Pretty.str " =>", Pretty.brk 1,  | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
340  | 
p, Pretty.str ")", Pretty.brk 1, parens (Pretty.block ps)]  | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
341  | 
else ps  | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
342  | 
end;  | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
343  | 
|
| 17521 | 344  | 
fun mk_v ((names, vs), s) = (case AList.lookup (op =) vs s of  | 
| 15531 | 345  | 
NONE => ((names, (s, [s])::vs), s)  | 
| 
20071
 
8f3e1ddb50e6
replaced Term.variant(list) by Name.variant(_list);
 
wenzelm 
parents: 
19861 
diff
changeset
 | 
346  | 
| SOME xs => let val s' = Name.variant names s in  | 
| 17521 | 347  | 
((s'::names, AList.update (op =) (s, s'::xs) vs), s') end);  | 
| 11537 | 348  | 
|
349  | 
fun distinct_v (nvs, Var ((s, 0), T)) =  | 
|
350  | 
apsnd (Var o rpair T o rpair 0) (mk_v (nvs, s))  | 
|
351  | 
| distinct_v (nvs, t $ u) =  | 
|
352  | 
let  | 
|
353  | 
val (nvs', t') = distinct_v (nvs, t);  | 
|
354  | 
val (nvs'', u') = distinct_v (nvs', u);  | 
|
355  | 
in (nvs'', t' $ u') end  | 
|
356  | 
| distinct_v x = x;  | 
|
357  | 
||
| 
15061
 
61a52739cd82
Added simple check that allows code generator to produce code containing
 
berghofe 
parents: 
15031 
diff
changeset
 | 
358  | 
fun is_exhaustive (Var _) = true  | 
| 
 
61a52739cd82
Added simple check that allows code generator to produce code containing
 
berghofe 
parents: 
15031 
diff
changeset
 | 
359  | 
  | is_exhaustive (Const ("Pair", _) $ t $ u) =
 | 
| 
 
61a52739cd82
Added simple check that allows code generator to produce code containing
 
berghofe 
parents: 
15031 
diff
changeset
 | 
360  | 
is_exhaustive t andalso is_exhaustive u  | 
| 
 
61a52739cd82
Added simple check that allows code generator to produce code containing
 
berghofe 
parents: 
15031 
diff
changeset
 | 
361  | 
| is_exhaustive _ = false;  | 
| 
 
61a52739cd82
Added simple check that allows code generator to produce code containing
 
berghofe 
parents: 
15031 
diff
changeset
 | 
362  | 
|
| 
 
61a52739cd82
Added simple check that allows code generator to produce code containing
 
berghofe 
parents: 
15031 
diff
changeset
 | 
363  | 
fun compile_match nvs eq_ps out_ps success_p can_fail =  | 
| 15570 | 364  | 
let val eqs = List.concat (separate [Pretty.str " andalso", Pretty.brk 1]  | 
365  | 
(map single (List.concat (map (mk_eq o snd) nvs) @ eq_ps)));  | 
|
| 11537 | 366  | 
in  | 
367  | 
Pretty.block  | 
|
368  | 
([Pretty.str "(fn ", mk_tuple out_ps, Pretty.str " =>", Pretty.brk 1] @  | 
|
369  | 
(Pretty.block ((if eqs=[] then [] else Pretty.str "if " ::  | 
|
370  | 
[Pretty.block eqs, Pretty.brk 1, Pretty.str "then "]) @  | 
|
371  | 
(success_p ::  | 
|
| 
15061
 
61a52739cd82
Added simple check that allows code generator to produce code containing
 
berghofe 
parents: 
15031 
diff
changeset
 | 
372  | 
(if eqs=[] then [] else [Pretty.brk 1, Pretty.str "else Seq.empty"]))) ::  | 
| 
 
61a52739cd82
Added simple check that allows code generator to produce code containing
 
berghofe 
parents: 
15031 
diff
changeset
 | 
373  | 
(if can_fail then  | 
| 
 
61a52739cd82
Added simple check that allows code generator to produce code containing
 
berghofe 
parents: 
15031 
diff
changeset
 | 
374  | 
[Pretty.brk 1, Pretty.str "| _ => Seq.empty)"]  | 
| 
 
61a52739cd82
Added simple check that allows code generator to produce code containing
 
berghofe 
parents: 
15031 
diff
changeset
 | 
375  | 
else [Pretty.str ")"])))  | 
| 11537 | 376  | 
end;  | 
377  | 
||
| 17144 | 378  | 
fun modename module s (iss, is) gr =  | 
379  | 
  let val (gr', id) = if s = "op =" then (gr, ("", "equal"))
 | 
|
380  | 
else mk_const_id module s gr  | 
|
381  | 
in (gr', space_implode "__"  | 
|
382  | 
(mk_qual_id module id ::  | 
|
383  | 
map (space_implode "_" o map string_of_int) (List.mapPartial I iss @ [is])))  | 
|
384  | 
end;  | 
|
| 11537 | 385  | 
|
| 17144 | 386  | 
fun compile_expr thy defs dep module brack (gr, (NONE, t)) =  | 
387  | 
apsnd single (invoke_codegen thy defs dep module brack (gr, t))  | 
|
388  | 
| compile_expr _ _ _ _ _ (gr, (SOME _, Var ((name, _), _))) =  | 
|
| 12557 | 389  | 
(gr, [Pretty.str name])  | 
| 17144 | 390  | 
| compile_expr thy defs dep module brack (gr, (SOME (Mode (mode, ms)), t)) =  | 
| 12557 | 391  | 
let  | 
392  | 
val (Const (name, _), args) = strip_comb t;  | 
|
| 17144 | 393  | 
val (gr', (ps, mode_id)) = foldl_map  | 
394  | 
(compile_expr thy defs dep module true) (gr, ms ~~ args) |>>>  | 
|
395  | 
modename module name mode;  | 
|
| 12557 | 396  | 
in (gr', (if brack andalso not (null ps) then  | 
397  | 
single o parens o Pretty.block else I)  | 
|
| 15570 | 398  | 
(List.concat (separate [Pretty.brk 1]  | 
| 17144 | 399  | 
([Pretty.str mode_id] :: ps))))  | 
| 12557 | 400  | 
end;  | 
401  | 
||
| 17144 | 402  | 
fun compile_clause thy defs gr dep module all_vs arg_vs modes (iss, is) (ts, ps) =  | 
| 11537 | 403  | 
let  | 
| 15570 | 404  | 
val modes' = modes @ List.mapPartial  | 
| 15531 | 405  | 
(fn (_, NONE) => NONE | (v, SOME js) => SOME (v, [([], js)]))  | 
| 12557 | 406  | 
(arg_vs ~~ iss);  | 
407  | 
||
| 11537 | 408  | 
fun check_constrt ((names, eqs), t) =  | 
409  | 
if is_constrt thy t then ((names, eqs), t) else  | 
|
| 
20071
 
8f3e1ddb50e6
replaced Term.variant(list) by Name.variant(_list);
 
wenzelm 
parents: 
19861 
diff
changeset
 | 
410  | 
let val s = Name.variant names "x";  | 
| 11537 | 411  | 
in ((s::names, (s, t)::eqs), Var ((s, 0), fastype_of t)) end;  | 
412  | 
||
| 
15031
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
413  | 
fun compile_eq (gr, (s, t)) =  | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
414  | 
apsnd (Pretty.block o cons (Pretty.str (s ^ " = ")) o single)  | 
| 17144 | 415  | 
(invoke_codegen thy defs dep module false (gr, t));  | 
| 
15031
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
416  | 
|
| 12557 | 417  | 
val (in_ts, out_ts) = get_args is 1 ts;  | 
| 11537 | 418  | 
val ((all_vs', eqs), in_ts') =  | 
419  | 
foldl_map check_constrt ((all_vs, []), in_ts);  | 
|
420  | 
||
| 
15031
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
421  | 
fun is_ind t = (case head_of t of  | 
| 17521 | 422  | 
Const (s, _) => s = "op =" orelse AList.defined (op =) modes s  | 
| 
15031
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
423  | 
| Var ((s, _), _) => s mem arg_vs);  | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
424  | 
|
| 11537 | 425  | 
fun compile_prems out_ts' vs names gr [] =  | 
426  | 
let  | 
|
| 12453 | 427  | 
val (gr2, out_ps) = foldl_map  | 
| 17144 | 428  | 
(invoke_codegen thy defs dep module false) (gr, out_ts);  | 
| 
15031
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
429  | 
val (gr3, eq_ps) = foldl_map compile_eq (gr2, eqs);  | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
430  | 
val ((names', eqs'), out_ts'') =  | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
431  | 
foldl_map check_constrt ((names, []), out_ts');  | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
432  | 
val (nvs, out_ts''') = foldl_map distinct_v  | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
433  | 
((names', map (fn x => (x, [x])) vs), out_ts'');  | 
| 12453 | 434  | 
val (gr4, out_ps') = foldl_map  | 
| 17144 | 435  | 
(invoke_codegen thy defs dep module false) (gr3, out_ts''');  | 
| 
15031
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
436  | 
val (gr5, eq_ps') = foldl_map compile_eq (gr4, eqs')  | 
| 11537 | 437  | 
in  | 
| 
15031
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
438  | 
(gr5, compile_match (snd nvs) (eq_ps @ eq_ps') out_ps'  | 
| 11537 | 439  | 
(Pretty.block [Pretty.str "Seq.single", Pretty.brk 1, mk_tuple out_ps])  | 
| 
15061
 
61a52739cd82
Added simple check that allows code generator to produce code containing
 
berghofe 
parents: 
15031 
diff
changeset
 | 
440  | 
(exists (not o is_exhaustive) out_ts'''))  | 
| 11537 | 441  | 
end  | 
442  | 
| compile_prems out_ts vs names gr ps =  | 
|
443  | 
let  | 
|
| 
19046
 
bc5c6c9b114e
removed distinct, renamed gen_distinct to distinct;
 
wenzelm 
parents: 
19025 
diff
changeset
 | 
444  | 
val vs' = distinct (op =) (List.concat (vs :: map term_vs out_ts));  | 
| 15531 | 445  | 
val SOME (p, mode as SOME (Mode ((_, js), _))) =  | 
| 
15126
 
54ae6c6ef3b1
Fixed bug in compile_clause that caused equality constraints
 
berghofe 
parents: 
15061 
diff
changeset
 | 
446  | 
select_mode_prem thy modes' vs' ps;  | 
| 11537 | 447  | 
val ps' = filter_out (equal p) ps;  | 
| 
15031
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
448  | 
val ((names', eqs), out_ts') =  | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
449  | 
foldl_map check_constrt ((names, []), out_ts);  | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
450  | 
val (nvs, out_ts'') = foldl_map distinct_v  | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
451  | 
((names', map (fn x => (x, [x])) vs), out_ts');  | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
452  | 
val (gr0, out_ps) = foldl_map  | 
| 17144 | 453  | 
(invoke_codegen thy defs dep module false) (gr, out_ts'');  | 
| 
15031
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
454  | 
val (gr1, eq_ps) = foldl_map compile_eq (gr0, eqs)  | 
| 11537 | 455  | 
in  | 
456  | 
(case p of  | 
|
| 12557 | 457  | 
Prem (us, t) =>  | 
| 11537 | 458  | 
let  | 
| 
15031
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
459  | 
val (in_ts, out_ts''') = get_args js 1 us;  | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
460  | 
val (gr2, in_ps) = foldl_map  | 
| 17144 | 461  | 
(invoke_codegen thy defs dep module false) (gr1, in_ts);  | 
| 
15031
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
462  | 
val (gr3, ps) = if is_ind t then  | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
463  | 
apsnd (fn ps => ps @ [Pretty.brk 1, mk_tuple in_ps])  | 
| 17144 | 464  | 
(compile_expr thy defs dep module false  | 
| 16645 | 465  | 
(gr2, (mode, t)))  | 
| 
15031
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
466  | 
else  | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
467  | 
apsnd (fn p => conv_ntuple us t  | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
468  | 
[Pretty.str "Seq.of_list", Pretty.brk 1, p])  | 
| 17144 | 469  | 
(invoke_codegen thy defs dep module true (gr2, t));  | 
| 
15031
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
470  | 
val (gr4, rest) = compile_prems out_ts''' vs' (fst nvs) gr3 ps';  | 
| 11537 | 471  | 
in  | 
| 
15031
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
472  | 
(gr4, compile_match (snd nvs) eq_ps out_ps  | 
| 12557 | 473  | 
(Pretty.block (ps @  | 
| 
15031
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
474  | 
[Pretty.str " :->", Pretty.brk 1, rest]))  | 
| 
15061
 
61a52739cd82
Added simple check that allows code generator to produce code containing
 
berghofe 
parents: 
15031 
diff
changeset
 | 
475  | 
(exists (not o is_exhaustive) out_ts''))  | 
| 11537 | 476  | 
end  | 
477  | 
| Sidecond t =>  | 
|
478  | 
let  | 
|
| 17144 | 479  | 
val (gr2, side_p) = invoke_codegen thy defs dep module true (gr1, t);  | 
| 11537 | 480  | 
val (gr3, rest) = compile_prems [] vs' (fst nvs) gr2 ps';  | 
481  | 
in  | 
|
| 
15031
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
482  | 
(gr3, compile_match (snd nvs) eq_ps out_ps  | 
| 11537 | 483  | 
(Pretty.block [Pretty.str "?? ", side_p,  | 
484  | 
Pretty.str " :->", Pretty.brk 1, rest])  | 
|
| 
15061
 
61a52739cd82
Added simple check that allows code generator to produce code containing
 
berghofe 
parents: 
15031 
diff
changeset
 | 
485  | 
(exists (not o is_exhaustive) out_ts''))  | 
| 11537 | 486  | 
end)  | 
487  | 
end;  | 
|
488  | 
||
| 
15126
 
54ae6c6ef3b1
Fixed bug in compile_clause that caused equality constraints
 
berghofe 
parents: 
15061 
diff
changeset
 | 
489  | 
val (gr', prem_p) = compile_prems in_ts' arg_vs all_vs' gr ps;  | 
| 11537 | 490  | 
in  | 
491  | 
(gr', Pretty.block [Pretty.str "Seq.single inp :->", Pretty.brk 1, prem_p])  | 
|
492  | 
end;  | 
|
493  | 
||
| 17144 | 494  | 
fun compile_pred thy defs gr dep module prfx all_vs arg_vs modes s cls mode =  | 
495  | 
let val (gr', (cl_ps, mode_id)) =  | 
|
496  | 
foldl_map (fn (gr, cl) => compile_clause thy defs  | 
|
497  | 
gr dep module all_vs arg_vs modes mode cl) (gr, cls) |>>>  | 
|
498  | 
modename module s mode  | 
|
| 11537 | 499  | 
in  | 
500  | 
((gr', "and "), Pretty.block  | 
|
501  | 
([Pretty.block (separate (Pretty.brk 1)  | 
|
| 17144 | 502  | 
(Pretty.str (prfx ^ mode_id) ::  | 
| 16645 | 503  | 
map Pretty.str arg_vs) @  | 
| 11537 | 504  | 
[Pretty.str " inp ="]),  | 
505  | 
Pretty.brk 1] @  | 
|
| 15570 | 506  | 
List.concat (separate [Pretty.str " ++", Pretty.brk 1] (map single cl_ps))))  | 
| 11537 | 507  | 
end;  | 
508  | 
||
| 17144 | 509  | 
fun compile_preds thy defs gr dep module all_vs arg_vs modes preds =  | 
| 11537 | 510  | 
let val ((gr', _), prs) = foldl_map (fn ((gr, prfx), (s, cls)) =>  | 
| 16645 | 511  | 
foldl_map (fn ((gr', prfx'), mode) => compile_pred thy defs gr'  | 
| 17144 | 512  | 
dep module prfx' all_vs arg_vs modes s cls mode)  | 
| 17521 | 513  | 
((gr, prfx), ((the o AList.lookup (op =) modes) s))) ((gr, "fun "), preds)  | 
| 11537 | 514  | 
in  | 
| 15570 | 515  | 
(gr', space_implode "\n\n" (map Pretty.string_of (List.concat prs)) ^ ";\n\n")  | 
| 11537 | 516  | 
end;  | 
517  | 
||
518  | 
(**** processing of introduction rules ****)  | 
|
519  | 
||
| 12557 | 520  | 
exception Modes of  | 
521  | 
(string * (int list option list * int list) list) list *  | 
|
| 17144 | 522  | 
(string * (int list list option list * int list list)) list;  | 
| 12557 | 523  | 
|
| 17144 | 524  | 
fun lookup_modes gr dep = apfst List.concat (apsnd List.concat (ListPair.unzip  | 
525  | 
(map ((fn (SOME (Modes x), _, _) => x | _ => ([], [])) o get_node gr)  | 
|
526  | 
(Graph.all_preds (fst gr) [dep]))));  | 
|
| 12557 | 527  | 
|
| 11537 | 528  | 
fun print_factors factors = message ("Factors:\n" ^
 | 
| 12557 | 529  | 
space_implode "\n" (map (fn (s, (fs, f)) => s ^ ": " ^  | 
530  | 
space_implode " -> " (map  | 
|
| 15531 | 531  | 
(fn NONE => "X" | SOME f' => string_of_factors [] f')  | 
532  | 
(fs @ [SOME f]))) factors));  | 
|
| 11537 | 533  | 
|
| 
15031
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
534  | 
fun prep_intrs intrs = map (rename_term o #prop o rep_thm o standard) intrs;  | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
535  | 
|
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
536  | 
fun constrain cs [] = []  | 
| 17521 | 537  | 
| constrain cs ((s, xs) :: ys) = (s, case AList.lookup (op =) cs s of  | 
| 15531 | 538  | 
NONE => xs  | 
539  | 
| SOME xs' => xs inter xs') :: constrain cs ys;  | 
|
| 
15031
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
540  | 
|
| 17144 | 541  | 
fun mk_extra_defs thy defs gr dep names module ts =  | 
| 15570 | 542  | 
Library.foldl (fn (gr, name) =>  | 
| 12557 | 543  | 
if name mem names then gr  | 
544  | 
else (case get_clauses thy name of  | 
|
| 15531 | 545  | 
NONE => gr  | 
| 16645 | 546  | 
| SOME (names, thyname, intrs) =>  | 
| 17144 | 547  | 
mk_ind_def thy defs gr dep names (if_library thyname module)  | 
548  | 
[] [] (prep_intrs intrs)))  | 
|
| 
15574
 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 
skalberg 
parents: 
15570 
diff
changeset
 | 
549  | 
(gr, foldr add_term_consts [] ts)  | 
| 12557 | 550  | 
|
| 17144 | 551  | 
and mk_ind_def thy defs gr dep names module modecs factorcs intrs =  | 
552  | 
add_edge (hd names, dep) gr handle Graph.UNDEF _ =>  | 
|
| 11537 | 553  | 
let  | 
| 
15031
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
554  | 
val _ $ (_ $ _ $ u) = Logic.strip_imp_concl (hd intrs);  | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
555  | 
val (_, args) = strip_comb u;  | 
| 15570 | 556  | 
val arg_vs = List.concat (map term_vs args);  | 
| 
15031
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
557  | 
|
| 
14250
 
d09e92e7c2bf
Inserted additional checks in functions dest_prem and add_prod_factors, to
 
berghofe 
parents: 
14195 
diff
changeset
 | 
558  | 
      fun dest_prem factors (_ $ (p as (Const ("op :", _) $ t $ u))) =
 | 
| 17521 | 559  | 
(case AList.lookup (op =) factors (case head_of u of  | 
| 
15031
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
560  | 
Const (name, _) => name | Var ((name, _), _) => name) of  | 
| 15531 | 561  | 
NONE => Prem (full_split_prod t, u)  | 
562  | 
| SOME f => Prem (split_prod [] f t, u))  | 
|
| 12557 | 563  | 
        | dest_prem factors (_ $ ((eq as Const ("op =", _)) $ t $ u)) =
 | 
564  | 
Prem ([t, u], eq)  | 
|
565  | 
| dest_prem factors (_ $ t) = Sidecond t;  | 
|
| 11537 | 566  | 
|
| 12557 | 567  | 
fun add_clause factors (clauses, intr) =  | 
| 11537 | 568  | 
let  | 
569  | 
val _ $ (_ $ t $ u) = Logic.strip_imp_concl intr;  | 
|
| 12557 | 570  | 
val Const (name, _) = head_of u;  | 
571  | 
val prems = map (dest_prem factors) (Logic.strip_imp_prems intr);  | 
|
| 11537 | 572  | 
in  | 
| 17521 | 573  | 
AList.update (op =) (name, ((these o AList.lookup (op =) clauses) name) @  | 
574  | 
[(split_prod [] ((the o AList.lookup (op =) factors) name) t, prems)]) clauses  | 
|
| 11537 | 575  | 
end;  | 
576  | 
||
| 19861 | 577  | 
fun check_set (Const (s, _)) = s mem names orelse is_some (get_clauses thy s)  | 
| 
15031
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
578  | 
| check_set (Var ((s, _), _)) = s mem arg_vs  | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
579  | 
| check_set _ = false;  | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
580  | 
|
| 12557 | 581  | 
      fun add_prod_factors extra_fs (fs, _ $ (Const ("op :", _) $ t $ u)) =
 | 
| 
15031
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
582  | 
if check_set (head_of u)  | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
583  | 
then infer_factors (sign_of thy) extra_fs  | 
| 15531 | 584  | 
(fs, (SOME (FVar (prod_factors [] t)), u))  | 
| 
15031
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
585  | 
else fs  | 
| 12557 | 586  | 
| add_prod_factors _ (fs, _) = fs;  | 
| 11537 | 587  | 
|
| 16645 | 588  | 
val gr' = mk_extra_defs thy defs  | 
| 17144 | 589  | 
(add_edge (hd names, dep)  | 
590  | 
(new_node (hd names, (NONE, "", "")) gr)) (hd names) names module intrs;  | 
|
591  | 
val (extra_modes, extra_factors) = lookup_modes gr' (hd names);  | 
|
| 
15031
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
592  | 
val fs = constrain factorcs (map (apsnd dest_factors)  | 
| 15570 | 593  | 
(Library.foldl (add_prod_factors extra_factors) ([], List.concat (map (fn t =>  | 
| 
15031
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
594  | 
Logic.strip_imp_concl t :: Logic.strip_imp_prems t) intrs))));  | 
| 15570 | 595  | 
val factors = List.mapPartial (fn (name, f) =>  | 
| 15531 | 596  | 
if name mem arg_vs then NONE  | 
| 17521 | 597  | 
else SOME (name, (map (AList.lookup (op =) fs) arg_vs, f))) fs;  | 
| 12557 | 598  | 
val clauses =  | 
| 15570 | 599  | 
Library.foldl (add_clause (fs @ map (apsnd snd) extra_factors)) ([], intrs);  | 
| 
15031
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
600  | 
val modes = constrain modecs  | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
601  | 
(infer_modes thy extra_modes factors arg_vs clauses);  | 
| 12557 | 602  | 
val _ = print_factors factors;  | 
| 11537 | 603  | 
val _ = print_modes modes;  | 
| 17144 | 604  | 
val (gr'', s) = compile_preds thy defs gr' (hd names) module (terms_vs intrs)  | 
605  | 
arg_vs (modes @ extra_modes) clauses;  | 
|
| 11537 | 606  | 
in  | 
| 17144 | 607  | 
(map_node (hd names)  | 
608  | 
(K (SOME (Modes (modes, factors)), module, s)) gr'')  | 
|
| 16645 | 609  | 
end;  | 
| 11537 | 610  | 
|
| 17144 | 611  | 
fun find_mode gr dep s u modes is = (case find_first (fn Mode ((_, js), _) => is=js)  | 
| 15660 | 612  | 
(modes_of modes u handle Option => []) of  | 
| 17144 | 613  | 
NONE => codegen_error gr dep  | 
614  | 
       ("No such mode for " ^ s ^ ": " ^ string_of_mode ([], is))
 | 
|
| 
15031
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
615  | 
| mode => mode);  | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
616  | 
|
| 17144 | 617  | 
fun mk_ind_call thy defs gr dep module t u is_query = (case head_of u of  | 
| 13038 | 618  | 
Const (s, T) => (case (get_clauses thy s, get_assoc_code thy s T) of  | 
| 15531 | 619  | 
(NONE, _) => NONE  | 
| 17144 | 620  | 
| (SOME (names, thyname, intrs), NONE) =>  | 
| 11537 | 621  | 
let  | 
| 12565 | 622  | 
          fun mk_mode (((ts, mode), i), Const ("dummy_pattern", _)) =
 | 
623  | 
((ts, mode), i+1)  | 
|
| 11537 | 624  | 
| mk_mode (((ts, mode), i), t) = ((ts @ [t], mode @ [i]), i+1);  | 
625  | 
||
| 17144 | 626  | 
val module' = if_library thyname module;  | 
| 16645 | 627  | 
val gr1 = mk_extra_defs thy defs  | 
| 17144 | 628  | 
(mk_ind_def thy defs gr dep names module'  | 
629  | 
[] [] (prep_intrs intrs)) dep names module' [u];  | 
|
630  | 
val (modes, factors) = lookup_modes gr1 dep;  | 
|
| 17521 | 631  | 
val ts = split_prod [] ((snd o the o AList.lookup (op =) factors) s) t;  | 
| 12557 | 632  | 
val (ts', is) = if is_query then  | 
| 15570 | 633  | 
fst (Library.foldl mk_mode ((([], []), 1), ts))  | 
| 11537 | 634  | 
else (ts, 1 upto length ts);  | 
| 17144 | 635  | 
val mode = find_mode gr1 dep s u modes is;  | 
| 12453 | 636  | 
val (gr2, in_ps) = foldl_map  | 
| 17144 | 637  | 
(invoke_codegen thy defs dep module false) (gr1, ts');  | 
| 16645 | 638  | 
val (gr3, ps) =  | 
| 17144 | 639  | 
compile_expr thy defs dep module false (gr2, (mode, u))  | 
| 11537 | 640  | 
in  | 
| 15531 | 641  | 
SOME (gr3, Pretty.block  | 
| 12557 | 642  | 
(ps @ [Pretty.brk 1, mk_tuple in_ps]))  | 
| 13038 | 643  | 
end  | 
| 15531 | 644  | 
| _ => NONE)  | 
645  | 
| _ => NONE);  | 
|
| 11537 | 646  | 
|
| 17144 | 647  | 
fun list_of_indset thy defs gr dep module brack u = (case head_of u of  | 
| 
15031
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
648  | 
Const (s, T) => (case (get_clauses thy s, get_assoc_code thy s T) of  | 
| 15531 | 649  | 
(NONE, _) => NONE  | 
| 17144 | 650  | 
| (SOME (names, thyname, intrs), NONE) =>  | 
| 
15031
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
651  | 
let  | 
| 17144 | 652  | 
val module' = if_library thyname module;  | 
| 16645 | 653  | 
val gr1 = mk_extra_defs thy defs  | 
| 17144 | 654  | 
(mk_ind_def thy defs gr dep names module'  | 
655  | 
[] [] (prep_intrs intrs)) dep names module' [u];  | 
|
656  | 
val (modes, factors) = lookup_modes gr1 dep;  | 
|
657  | 
val mode = find_mode gr1 dep s u modes [];  | 
|
| 16645 | 658  | 
val (gr2, ps) =  | 
| 
18388
 
ab1a710a68ce
list_of_indset now also generates code for set type.
 
berghofe 
parents: 
17521 
diff
changeset
 | 
659  | 
compile_expr thy defs dep module false (gr1, (mode, u));  | 
| 
 
ab1a710a68ce
list_of_indset now also generates code for set type.
 
berghofe 
parents: 
17521 
diff
changeset
 | 
660  | 
val (gr3, _) =  | 
| 
 
ab1a710a68ce
list_of_indset now also generates code for set type.
 
berghofe 
parents: 
17521 
diff
changeset
 | 
661  | 
invoke_tycodegen thy defs dep module false (gr2, body_type T)  | 
| 
15031
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
662  | 
in  | 
| 
18388
 
ab1a710a68ce
list_of_indset now also generates code for set type.
 
berghofe 
parents: 
17521 
diff
changeset
 | 
663  | 
SOME (gr3, (if brack then parens else I)  | 
| 
15031
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
664  | 
(Pretty.block ([Pretty.str "Seq.list_of", Pretty.brk 1,  | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
665  | 
               Pretty.str "("] @
 | 
| 17521 | 666  | 
conv_ntuple' (snd (valOf (AList.lookup (op =) factors s)))  | 
| 
15031
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
667  | 
(HOLogic.dest_setT (fastype_of u))  | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
668  | 
(ps @ [Pretty.brk 1, Pretty.str "()"]) @  | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
669  | 
[Pretty.str ")"])))  | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
670  | 
end  | 
| 15531 | 671  | 
| _ => NONE)  | 
672  | 
| _ => NONE);  | 
|
| 
15031
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
673  | 
|
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
674  | 
fun clause_of_eqn eqn =  | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
675  | 
let  | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
676  | 
val (t, u) = HOLogic.dest_eq (HOLogic.dest_Trueprop (concl_of eqn));  | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
677  | 
val (Const (s, T), ts) = strip_comb t;  | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
678  | 
val (Ts, U) = strip_type T  | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
679  | 
in  | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
680  | 
rename_term  | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
681  | 
(Logic.list_implies (prems_of eqn, HOLogic.mk_Trueprop (HOLogic.mk_mem  | 
| 17144 | 682  | 
(foldr1 HOLogic.mk_prod (ts @ [u]), Const (s ^ "_aux",  | 
| 
15031
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
683  | 
HOLogic.mk_setT (foldr1 HOLogic.mk_prodT (Ts @ [U])))))))  | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
684  | 
end;  | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
685  | 
|
| 17144 | 686  | 
fun mk_fun thy defs name eqns dep module module' gr =  | 
687  | 
case try (get_node gr) name of  | 
|
688  | 
NONE =>  | 
|
| 
15031
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
689  | 
let  | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
690  | 
val clauses = map clause_of_eqn eqns;  | 
| 17144 | 691  | 
val pname = name ^ "_aux";  | 
| 
15031
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
692  | 
val arity = length (snd (strip_comb (fst (HOLogic.dest_eq  | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
693  | 
(HOLogic.dest_Trueprop (concl_of (hd eqns)))))));  | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
694  | 
val mode = 1 upto arity;  | 
| 17144 | 695  | 
val (gr', (fun_id, mode_id)) = gr |>  | 
696  | 
mk_const_id module' name |>>>  | 
|
697  | 
modename module' pname ([], mode);  | 
|
| 
15031
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
698  | 
      val vars = map (fn i => Pretty.str ("x" ^ string_of_int i)) mode;
 | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
699  | 
val s = Pretty.string_of (Pretty.block  | 
| 17144 | 700  | 
        [mk_app false (Pretty.str ("fun " ^ snd fun_id)) vars, Pretty.str " =",
 | 
| 
15031
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
701  | 
Pretty.brk 1, Pretty.str "Seq.hd", Pretty.brk 1,  | 
| 17144 | 702  | 
parens (Pretty.block [Pretty.str mode_id,  | 
| 
15031
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
703  | 
Pretty.brk 1, mk_tuple vars])]) ^ ";\n\n";  | 
| 17144 | 704  | 
val gr'' = mk_ind_def thy defs (add_edge (name, dep)  | 
705  | 
(new_node (name, (NONE, module', s)) gr')) name [pname] module'  | 
|
| 
15031
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
706  | 
[(pname, [([], mode)])]  | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
707  | 
[(pname, map (fn i => replicate i 2) (0 upto arity-1))]  | 
| 
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
708  | 
clauses;  | 
| 17144 | 709  | 
val (modes, _) = lookup_modes gr'' dep;  | 
710  | 
val _ = find_mode gr'' dep pname (snd (HOLogic.dest_mem (HOLogic.dest_Trueprop  | 
|
| 
15031
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
711  | 
(Logic.strip_imp_concl (hd clauses))))) modes mode  | 
| 17144 | 712  | 
in (gr'', mk_qual_id module fun_id) end  | 
713  | 
| SOME _ =>  | 
|
714  | 
(add_edge (name, dep) gr, mk_qual_id module (get_const_id name gr));  | 
|
| 
15031
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
715  | 
|
| 17144 | 716  | 
fun inductive_codegen thy defs gr dep module brack (Const ("op :", _) $ t $ u) =
 | 
717  | 
((case mk_ind_call thy defs gr dep module (Term.no_dummy_patterns t) u false of  | 
|
| 15531 | 718  | 
NONE => NONE  | 
719  | 
| SOME (gr', call_p) => SOME (gr', (if brack then parens else I)  | 
|
| 12453 | 720  | 
           (Pretty.block [Pretty.str "?! (", call_p, Pretty.str ")"])))
 | 
| 17144 | 721  | 
handle TERM _ => mk_ind_call thy defs gr dep module t u true)  | 
722  | 
| inductive_codegen thy defs gr dep module brack t = (case strip_comb t of  | 
|
| 17412 | 723  | 
(Const (s, _), ts) => (case Symtab.lookup (#eqns (CodegenData.get thy)) s of  | 
| 17144 | 724  | 
NONE => list_of_indset thy defs gr dep module brack t  | 
| 15531 | 725  | 
| SOME eqns =>  | 
| 
15031
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
726  | 
let  | 
| 17144 | 727  | 
val (_, (_, thyname)) = split_last eqns;  | 
| 16645 | 728  | 
val (gr', id) = mk_fun thy defs s (preprocess thy (map fst eqns))  | 
| 17144 | 729  | 
dep module (if_library thyname module) gr;  | 
| 16645 | 730  | 
val (gr'', ps) = foldl_map  | 
| 17144 | 731  | 
(invoke_codegen thy defs dep module true) (gr', ts);  | 
| 16645 | 732  | 
in SOME (gr'', mk_app brack (Pretty.str id) ps)  | 
| 
15031
 
726dc9146bb1
- Added support for conditional equations whose premises involve
 
berghofe 
parents: 
14981 
diff
changeset
 | 
733  | 
end)  | 
| 15531 | 734  | 
| _ => NONE);  | 
| 11537 | 735  | 
|
| 12557 | 736  | 
val setup =  | 
| 18708 | 737  | 
add_codegen "inductive" inductive_codegen #>  | 
738  | 
CodegenData.init #>  | 
|
739  | 
add_attribute "ind" (Scan.option (Args.$$$ "target" |-- Args.colon |-- Args.name) >> add);  | 
|
| 11537 | 740  | 
|
741  | 
end;  | 
|
| 12453 | 742  | 
|
743  | 
||
744  | 
(**** combinators for code generated from inductive predicates ****)  | 
|
745  | 
||
746  | 
infix 5 :->;  | 
|
747  | 
infix 3 ++;  | 
|
748  | 
||
| 19861 | 749  | 
fun s :-> f = Seq.maps f s;  | 
| 12453 | 750  | 
|
| 19861 | 751  | 
fun s1 ++ s2 = Seq.append s1 s2;  | 
| 12453 | 752  | 
|
753  | 
fun ?? b = if b then Seq.single () else Seq.empty;  | 
|
754  | 
||
| 19861 | 755  | 
fun ?! s = is_some (Seq.pull s);  | 
| 12453 | 756  | 
|
| 17144 | 757  | 
fun equal__1 x = Seq.single x;  | 
| 12453 | 758  | 
|
| 17144 | 759  | 
val equal__2 = equal__1;  | 
| 12557 | 760  | 
|
| 17144 | 761  | 
fun equal__1_2 (x, y) = ?? (x = y);  |